From f9776b406755f68d57b24bac411f6a80c9882cb0 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Mon, 17 Apr 2023 10:44:09 +0000 Subject: [PATCH] bump to nightly-2023-04-17-10 mathlib commit https://github.com/leanprover-community/mathlib/commit/17ad94b4953419f3e3ce3e77da3239c62d1d09f0 --- .../AlgebraicTopology/TopologicalSimplex.lean | 32 +++++ Mathbin/Analysis/Convex/Exposed.lean | 112 +++++++++++++++ Mathbin/Analysis/Normed/Ring/Seminorm.lean | 130 +++++++++++++++++- Mathbin/CategoryTheory/Sites/Whiskering.lean | 48 +++++++ Mathbin/Data/Int/Cast/Lemmas.lean | 2 +- Mathbin/GroupTheory/FreeAbelianGroup.lean | 22 +-- .../GroupTheory/FreeAbelianGroupFinsupp.lean | 124 +++++++++++++++++ Mathbin/GroupTheory/OrderOfElement.lean | 8 +- Mathbin/Order/Category/BddDistLat.lean | 17 +-- Mathbin/Order/Category/DistLat.lean | 90 ++++++++---- Mathbin/Order/Filter/Pointwise.lean | 16 +-- Mathbin/Topology/Algebra/Group/Basic.lean | 34 ++--- Mathbin/Topology/List.lean | 20 +-- Mathbin/Topology/Separation.lean | 2 +- lake-manifest.json | 8 +- lakefile.lean | 4 +- 16 files changed, 570 insertions(+), 99 deletions(-) diff --git a/Mathbin/AlgebraicTopology/TopologicalSimplex.lean b/Mathbin/AlgebraicTopology/TopologicalSimplex.lean index 77fecba28a..a4143914e9 100644 --- a/Mathbin/AlgebraicTopology/TopologicalSimplex.lean +++ b/Mathbin/AlgebraicTopology/TopologicalSimplex.lean @@ -30,6 +30,12 @@ open Simplicial NNReal BigOperators Classical attribute [local instance] CategoryTheory.ConcreteCategory.hasCoeToSort CategoryTheory.ConcreteCategory.hasCoeToFun +/- warning: simplex_category.to_Top_obj -> SimplexCategory.toTopObj is a dubious translation: +lean 3 declaration is + forall (x : SimplexCategory), Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal) +but is expected to have type + forall (x : SimplexCategory), Set.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) +Case conversion may be inaccurate. Consider using '#align simplex_category.to_Top_obj SimplexCategory.toTopObjₓ'. -/ /-- The topological simplex associated to `x : simplex_category`. This is the object part of the functor `simplex_category.to_Top`. -/ def toTopObj (x : SimplexCategory) := @@ -39,11 +45,23 @@ def toTopObj (x : SimplexCategory) := instance (x : SimplexCategory) : CoeFun x.toTopObj fun _ => x → ℝ≥0 := ⟨fun f => (f : x → ℝ≥0)⟩ +/- warning: simplex_category.to_Top_obj.ext -> SimplexCategory.toTopObj.ext is a dubious translation: +lean 3 declaration is + forall {x : SimplexCategory} (f : coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) (g : coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)), (Eq.{1} ((fun (_x : coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) => (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal) f) (coeFn.{1, 1} (coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) (fun (_x : coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) => (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal) (SimplexCategory.toTopObj.hasCoeToFun x) f) (coeFn.{1, 1} (coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) (fun (_x : coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) => (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal) (SimplexCategory.toTopObj.hasCoeToFun x) g)) -> (Eq.{1} (coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) f g) +but is expected to have type + forall {x : SimplexCategory} (f : Set.Elem.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (SimplexCategory.toTopObj x)) (g : Set.Elem.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (SimplexCategory.toTopObj x)), (Eq.{1} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (Subtype.val.{1} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (fun (x_1 : (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) => Membership.mem.{0, 0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (Set.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal)) (Set.instMembershipSet.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal)) x_1 (SimplexCategory.toTopObj x)) f) (Subtype.val.{1} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (fun (x_1 : (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) => Membership.mem.{0, 0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (Set.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal)) (Set.instMembershipSet.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal)) x_1 (SimplexCategory.toTopObj x)) g)) -> (Eq.{1} (Set.Elem.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (SimplexCategory.toTopObj x)) f g) +Case conversion may be inaccurate. Consider using '#align simplex_category.to_Top_obj.ext SimplexCategory.toTopObj.extₓ'. -/ @[ext] theorem toTopObj.ext {x : SimplexCategory} (f g : x.toTopObj) : (f : x → ℝ≥0) = g → f = g := Subtype.ext #align simplex_category.to_Top_obj.ext SimplexCategory.toTopObj.ext +/- warning: simplex_category.to_Top_map -> SimplexCategory.toTopMap is a dubious translation: +lean 3 declaration is + forall {x : SimplexCategory} {y : SimplexCategory}, (Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y) -> (coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) -> (coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) (SimplexCategory.toTopObj y)) +but is expected to have type + forall {x : SimplexCategory} {y : SimplexCategory}, (Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y) -> (Set.Elem.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (SimplexCategory.toTopObj x)) -> (Set.Elem.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal) (SimplexCategory.toTopObj y)) +Case conversion may be inaccurate. Consider using '#align simplex_category.to_Top_map SimplexCategory.toTopMapₓ'. -/ /-- A morphism in `simplex_category` induces a map on the associated topological spaces. -/ def toTopMap {x y : SimplexCategory} (f : x ⟶ y) : x.toTopObj → y.toTopObj := fun g => ⟨fun i => ∑ j in Finset.univ.filterₓ fun k => f k = i, g j, @@ -64,12 +82,24 @@ def toTopMap {x y : SimplexCategory} (f : x ⟶ y) : x.toTopObj → y.toTopObj : rw [← he.1, ← he.2]⟩ #align simplex_category.to_Top_map SimplexCategory.toTopMap +/- warning: simplex_category.coe_to_Top_map -> SimplexCategory.coe_toTopMap is a dubious translation: +lean 3 declaration is + forall {x : SimplexCategory} {y : SimplexCategory} (f : Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y) (g : coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) (i : coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y), Eq.{1} NNReal (coeFn.{1, 1} (coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) (SimplexCategory.toTopObj y)) (fun (_x : coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) (SimplexCategory.toTopObj y)) => (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal) (SimplexCategory.toTopObj.hasCoeToFun y) (SimplexCategory.toTopMap x y f g) i) (Finset.sum.{0, 0} NNReal (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) (NonUnitalNonAssocSemiring.toAddCommMonoid.{0} NNReal (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} NNReal (Semiring.toNonAssocSemiring.{0} NNReal NNReal.semiring))) (Finset.filter.{0} (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) (fun (k : coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) => Eq.{1} (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) (coeFn.{1, 1} (Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y) (fun (f : Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y) => (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y)) (CategoryTheory.ConcreteCategory.hasCoeToFun.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory x y) f k) i) (fun (a : coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) => Fin.decidableEq (HAdd.hAdd.{0, 0, 0} Nat Nat Nat (instHAdd.{0} Nat Nat.hasAdd) (SimplexCategory.len y) (OfNat.ofNat.{0} Nat 1 (OfNat.mk.{0} Nat 1 (One.one.{0} Nat Nat.hasOne)))) (coeFn.{1, 1} (Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y) (fun (f : Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y) => (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y)) (CategoryTheory.ConcreteCategory.hasCoeToFun.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory x y) f a) i) (Finset.univ.{0} (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) (Fin.fintype (HAdd.hAdd.{0, 0, 0} Nat Nat Nat (instHAdd.{0} Nat Nat.hasAdd) (SimplexCategory.len x) (OfNat.ofNat.{0} Nat 1 (OfNat.mk.{0} Nat 1 (One.one.{0} Nat Nat.hasOne))))))) (fun (j : coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) => coeFn.{1, 1} (coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) (fun (_x : coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) => (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal) (SimplexCategory.toTopObj.hasCoeToFun x) g j)) +but is expected to have type + forall {x : SimplexCategory} {y : SimplexCategory} (f : Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y) (g : Set.Elem.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (SimplexCategory.toTopObj x)) (i : Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y), Eq.{1} NNReal (Subtype.val.{1} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal) (fun (x : (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal) => Membership.mem.{0, 0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal) (Set.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal)) (Set.instMembershipSet.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal)) x (SimplexCategory.toTopObj y)) (SimplexCategory.toTopMap x y f g) i) (Finset.sum.{0, 0} NNReal (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) (NonUnitalNonAssocSemiring.toAddCommMonoid.{0} NNReal (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} NNReal (Semiring.toNonAssocSemiring.{0} NNReal instNNRealSemiring))) (Finset.filter.{0} (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) (fun (k : Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) => Eq.{1} (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) (Prefunctor.map.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x y f k) i) (fun (a : Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) => Classical.propDecidable ((fun (k : Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) => Eq.{1} (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) (Prefunctor.map.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x y f k) i) a)) (Finset.univ.{0} (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) (SimplexCategory.instFintypeObjSimplexCategoryToQuiverToCategoryStructSmallCategoryTypeToQuiverToCategoryStructTypesToPrefunctorForget x))) (fun (j : Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) => Subtype.val.{1} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (fun (x_1 : (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) => Membership.mem.{0, 0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (Set.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal)) (Set.instMembershipSet.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal)) x_1 (SimplexCategory.toTopObj x)) g j)) +Case conversion may be inaccurate. Consider using '#align simplex_category.coe_to_Top_map SimplexCategory.coe_toTopMapₓ'. -/ @[simp] theorem coe_toTopMap {x y : SimplexCategory} (f : x ⟶ y) (g : x.toTopObj) (i : y) : toTopMap f g i = ∑ j in Finset.univ.filterₓ fun k => f k = i, g j := rfl #align simplex_category.coe_to_Top_map SimplexCategory.coe_toTopMap +/- warning: simplex_category.continuous_to_Top_map -> SimplexCategory.continuous_toTopMap is a dubious translation: +lean 3 declaration is + forall {x : SimplexCategory} {y : SimplexCategory} (f : Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y), Continuous.{0, 0} (coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (SimplexCategory.toTopObj x)) (coeSort.{1, 2} (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) Type (Set.hasCoeToSort.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) (SimplexCategory.toTopObj y)) (Subtype.topologicalSpace.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal) (fun (x_1 : (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal) => Membership.Mem.{0, 0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal) (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) (Set.hasMem.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) -> NNReal)) x_1 (SimplexCategory.toTopObj x)) (Pi.topologicalSpace.{0, 0} (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) (fun (ᾰ : coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) => NNReal) (fun (a : coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) x) => NNReal.topologicalSpace))) (Subtype.topologicalSpace.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal) (fun (x : (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal) => Membership.Mem.{0, 0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal) (Set.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) (Set.hasMem.{0} ((coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) -> NNReal)) x (SimplexCategory.toTopObj y)) (Pi.topologicalSpace.{0, 0} (coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) (fun (ᾰ : coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) => NNReal) (fun (a : coeSort.{1, 2} SimplexCategory Type (CategoryTheory.ConcreteCategory.hasCoeToSort.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.CategoryTheory.concreteCategory) y) => NNReal.topologicalSpace))) (SimplexCategory.toTopMap x y f) +but is expected to have type + forall {x : SimplexCategory} {y : SimplexCategory} (f : Quiver.Hom.{1, 0} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) x y), Continuous.{0, 0} (Set.Elem.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (SimplexCategory.toTopObj x)) (Set.Elem.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal) (SimplexCategory.toTopObj y)) (instTopologicalSpaceSubtype.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (fun (x_1 : (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) => Membership.mem.{0, 0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal) (Set.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal)) (Set.instMembershipSet.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) -> NNReal)) x_1 (SimplexCategory.toTopObj x)) (Pi.topologicalSpace.{0, 0} (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) (fun (ᾰ : Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) => NNReal) (fun (a : Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) x) => NNReal.instTopologicalSpaceNNReal))) (instTopologicalSpaceSubtype.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal) (fun (x : (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal) => Membership.mem.{0, 0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal) (Set.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal)) (Set.instMembershipSet.{0} ((Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) -> NNReal)) x (SimplexCategory.toTopObj y)) (Pi.topologicalSpace.{0, 0} (Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) (fun (ᾰ : Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) => NNReal) (fun (a : Prefunctor.obj.{1, 1, 0, 1} SimplexCategory (CategoryTheory.CategoryStruct.toQuiver.{0, 0} SimplexCategory (CategoryTheory.Category.toCategoryStruct.{0, 0} SimplexCategory SimplexCategory.smallCategory)) Type (CategoryTheory.CategoryStruct.toQuiver.{0, 1} Type (CategoryTheory.Category.toCategoryStruct.{0, 1} Type CategoryTheory.types.{0})) (CategoryTheory.Functor.toPrefunctor.{0, 0, 0, 1} SimplexCategory SimplexCategory.smallCategory Type CategoryTheory.types.{0} (CategoryTheory.ConcreteCategory.Forget.{0, 0, 0} SimplexCategory SimplexCategory.smallCategory SimplexCategory.instConcreteCategorySimplexCategorySmallCategory)) y) => NNReal.instTopologicalSpaceNNReal))) (SimplexCategory.toTopMap x y f) +Case conversion may be inaccurate. Consider using '#align simplex_category.continuous_to_Top_map SimplexCategory.continuous_toTopMapₓ'. -/ @[continuity] theorem continuous_toTopMap {x y : SimplexCategory} (f : x ⟶ y) : Continuous (toTopMap f) := Continuous.subtype_mk @@ -78,6 +108,7 @@ theorem continuous_toTopMap {x y : SimplexCategory} (f : x ⟶ y) : Continuous ( _ #align simplex_category.continuous_to_Top_map SimplexCategory.continuous_toTopMap +#print SimplexCategory.toTop /- /-- The functor associating the topological `n`-simplex to `[n] : simplex_category`. -/ @[simps] def toTop : SimplexCategory ⥤ TopCat @@ -105,6 +136,7 @@ def toTop : SimplexCategory ⥤ TopCat Finset.mem_inter] at he rw [← he.1, ← he.2] #align simplex_category.to_Top SimplexCategory.toTop +-/ end SimplexCategory diff --git a/Mathbin/Analysis/Convex/Exposed.lean b/Mathbin/Analysis/Convex/Exposed.lean index b94bf72d98..3ff06a891a 100644 --- a/Mathbin/Analysis/Convex/Exposed.lean +++ b/Mathbin/Analysis/Convex/Exposed.lean @@ -58,11 +58,13 @@ section PreorderSemiring variable (𝕜 : Type _) {E : Type _} [TopologicalSpace 𝕜] [Semiring 𝕜] [Preorder 𝕜] [AddCommMonoid E] [TopologicalSpace E] [Module 𝕜 E] {A B : Set E} +#print IsExposed /- /-- A set `B` is exposed with respect to `A` iff it maximizes some functional over `A` (and contains all points maximizing it). Written `is_exposed 𝕜 A B`. -/ def IsExposed (A B : Set E) : Prop := B.Nonempty → ∃ l : E →L[𝕜] 𝕜, B = { x ∈ A | ∀ y ∈ A, l y ≤ l x } #align is_exposed IsExposed +-/ end PreorderSemiring @@ -71,15 +73,29 @@ section OrderedRing variable {𝕜 : Type _} {E : Type _} [TopologicalSpace 𝕜] [OrderedRing 𝕜] [AddCommMonoid E] [TopologicalSpace E] [Module 𝕜 E] {l : E →L[𝕜] 𝕜} {A B C : Set E} {X : Finset E} {x : E} +#print ContinuousLinearMap.toExposed /- /-- A useful way to build exposed sets from intersecting `A` with halfspaces (modelled by an inequality with a functional). -/ def ContinuousLinearMap.toExposed (l : E →L[𝕜] 𝕜) (A : Set E) : Set E := { x ∈ A | ∀ y ∈ A, l y ≤ l x } #align continuous_linear_map.to_exposed ContinuousLinearMap.toExposed +-/ +/- warning: continuous_linear_map.to_exposed.is_exposed -> ContinuousLinearMap.toExposed.isExposed is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {l : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))} {A : Set.{u2} E}, IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A (ContinuousLinearMap.toExposed.{u1, u2} 𝕜 E _inst_1 _inst_2 _inst_3 _inst_4 _inst_5 l A) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] {l : ContinuousLinearMap.{u2, u2, u1, u2} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (RingHom.id.{u2} 𝕜 (NonAssocRing.toNonAssocSemiring.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u2} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u2} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u2} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2))} {A : Set.{u1} E}, IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A (ContinuousLinearMap.toExposed.{u2, u1} 𝕜 E _inst_1 _inst_2 _inst_3 _inst_4 _inst_5 l A) +Case conversion may be inaccurate. Consider using '#align continuous_linear_map.to_exposed.is_exposed ContinuousLinearMap.toExposed.isExposedₓ'. -/ theorem ContinuousLinearMap.toExposed.isExposed : IsExposed 𝕜 A (l.toExposed A) := fun h => ⟨l, rfl⟩ #align continuous_linear_map.to_exposed.is_exposed ContinuousLinearMap.toExposed.isExposed +/- warning: is_exposed_empty -> isExposed_empty is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E}, IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A (EmptyCollection.emptyCollection.{u2} (Set.{u2} E) (Set.hasEmptyc.{u2} E)) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] {A : Set.{u1} E}, IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A (EmptyCollection.emptyCollection.{u1} (Set.{u1} E) (Set.instEmptyCollectionSet.{u1} E)) +Case conversion may be inaccurate. Consider using '#align is_exposed_empty isExposed_emptyₓ'. -/ theorem isExposed_empty : IsExposed 𝕜 A ∅ := fun ⟨x, hx⟩ => by exfalso @@ -88,6 +104,12 @@ theorem isExposed_empty : IsExposed 𝕜 A ∅ := fun ⟨x, hx⟩ => namespace IsExposed +/- warning: is_exposed.subset -> IsExposed.subset is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} {B : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B) -> (HasSubset.Subset.{u2} (Set.{u2} E) (Set.hasSubset.{u2} E) B A) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] {A : Set.{u1} E} {B : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A B) -> (HasSubset.Subset.{u1} (Set.{u1} E) (Set.instHasSubsetSet.{u1} E) B A) +Case conversion may be inaccurate. Consider using '#align is_exposed.subset IsExposed.subsetₓ'. -/ protected theorem subset (hAB : IsExposed 𝕜 A B) : B ⊆ A := by rintro x hx @@ -95,15 +117,29 @@ protected theorem subset (hAB : IsExposed 𝕜 A B) : B ⊆ A := exact hx.1 #align is_exposed.subset IsExposed.subset +#print IsExposed.refl /- @[refl] protected theorem refl (A : Set E) : IsExposed 𝕜 A A := fun ⟨w, hw⟩ => ⟨0, Subset.antisymm (fun x hx => ⟨hx, fun y hy => le_refl 0⟩) fun x hx => hx.1⟩ #align is_exposed.refl IsExposed.refl +-/ +/- warning: is_exposed.antisymm -> IsExposed.antisymm is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} {B : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B) -> (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 B A) -> (Eq.{succ u2} (Set.{u2} E) A B) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] {A : Set.{u1} E} {B : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A B) -> (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 B A) -> (Eq.{succ u1} (Set.{u1} E) A B) +Case conversion may be inaccurate. Consider using '#align is_exposed.antisymm IsExposed.antisymmₓ'. -/ protected theorem antisymm (hB : IsExposed 𝕜 A B) (hA : IsExposed 𝕜 B A) : A = B := hA.Subset.antisymm hB.Subset #align is_exposed.antisymm IsExposed.antisymm +/- warning: is_exposed.mono -> IsExposed.mono is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} {B : Set.{u2} E} {C : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A C) -> (HasSubset.Subset.{u2} (Set.{u2} E) (Set.hasSubset.{u2} E) B A) -> (HasSubset.Subset.{u2} (Set.{u2} E) (Set.hasSubset.{u2} E) C B) -> (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 B C) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] {A : Set.{u1} E} {B : Set.{u1} E} {C : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A C) -> (HasSubset.Subset.{u1} (Set.{u1} E) (Set.instHasSubsetSet.{u1} E) B A) -> (HasSubset.Subset.{u1} (Set.{u1} E) (Set.instHasSubsetSet.{u1} E) C B) -> (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 B C) +Case conversion may be inaccurate. Consider using '#align is_exposed.mono IsExposed.monoₓ'. -/ /- `is_exposed` is *not* transitive: Consider a (topologically) open cube with vertices `A₀₀₀, ..., A₁₁₁` and add to it the triangle `A₀₀₀A₀₀₁A₀₁₀`. Then `A₀₀₁A₀₁₀` is an exposed subset of `A₀₀₀A₀₀₁A₀₁₀` which is an exposed subset of the cube, but `A₀₀₁A₀₁₀` is not itself an exposed @@ -118,6 +154,12 @@ protected theorem mono (hC : IsExposed 𝕜 A C) (hBA : B ⊆ A) (hCB : C ⊆ B) ⟨hBA hx.1, fun y hy => (hw.2 y hy).trans (hx.2 w (hCB hw))⟩⟩ #align is_exposed.mono IsExposed.mono +/- warning: is_exposed.eq_inter_halfspace' -> IsExposed.eq_inter_halfspace' is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} {B : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B) -> (Set.Nonempty.{u2} E B) -> (Exists.{max (succ u2) (succ u1)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) (fun (l : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) => Exists.{succ u1} 𝕜 (fun (a : 𝕜) => Eq.{succ u2} (Set.{u2} E) B (Sep.sep.{u2, u2} E (Set.{u2} E) (Set.hasSep.{u2} E) (fun (x : E) => LE.le.{u1} 𝕜 (Preorder.toLE.{u1} 𝕜 (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2)))) a (coeFn.{max (succ u2) (succ u1), max (succ u2) (succ u1)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) (fun (_x : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) => E -> 𝕜) (ContinuousLinearMap.toFun.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) l x)) A)))) +but is expected to have type + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} {B : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedRing.toPartialOrder.{u1} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A B) -> (Set.Nonempty.{u2} E B) -> (Exists.{max (succ u1) (succ u2)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) (fun (l : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) => Exists.{succ u1} 𝕜 (fun (a : 𝕜) => Eq.{succ u2} (Set.{u2} E) B (setOf.{u2} E (fun (x : E) => And (Membership.mem.{u2, u2} E (Set.{u2} E) (Set.instMembershipSet.{u2} E) x A) (LE.le.{u1} 𝕜 (Preorder.toLE.{u1} 𝕜 (PartialOrder.toPreorder.{u1} 𝕜 (OrderedRing.toPartialOrder.{u1} 𝕜 _inst_2))) a (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E (fun (a : E) => (fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) a) (ContinuousMapClass.toFunLike.{max u1 u2, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E 𝕜 _inst_4 _inst_1 (ContinuousSemilinearMapClass.toContinuousMapClass.{max u1 u2, u1, u1, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (ContinuousLinearMap.continuousSemilinearMapClass.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))))) l x))))))) +Case conversion may be inaccurate. Consider using '#align is_exposed.eq_inter_halfspace' IsExposed.eq_inter_halfspace'ₓ'. -/ /-- If `B` is a nonempty exposed subset of `A`, then `B` is the intersection of `A` with some closed halfspace. The converse is *not* true. It would require that the corresponding open halfspace doesn't intersect `A`. -/ @@ -132,6 +174,12 @@ theorem eq_inter_halfspace' {A B : Set E} (hAB : IsExposed 𝕜 A B) (hB : B.Non ⟨hx.1, fun y hy => (hw.2 y hy).trans hx.2⟩⟩ #align is_exposed.eq_inter_halfspace' IsExposed.eq_inter_halfspace' +/- warning: is_exposed.eq_inter_halfspace -> IsExposed.eq_inter_halfspace is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] [_inst_6 : Nontrivial.{u1} 𝕜] {A : Set.{u2} E} {B : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B) -> (Exists.{max (succ u2) (succ u1)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) (fun (l : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) => Exists.{succ u1} 𝕜 (fun (a : 𝕜) => Eq.{succ u2} (Set.{u2} E) B (Sep.sep.{u2, u2} E (Set.{u2} E) (Set.hasSep.{u2} E) (fun (x : E) => LE.le.{u1} 𝕜 (Preorder.toLE.{u1} 𝕜 (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2)))) a (coeFn.{max (succ u2) (succ u1), max (succ u2) (succ u1)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) (fun (_x : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) => E -> 𝕜) (ContinuousLinearMap.toFun.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) l x)) A)))) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] [_inst_6 : Nontrivial.{u2} 𝕜] {A : Set.{u1} E} {B : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A B) -> (Exists.{max (succ u2) (succ u1)} (ContinuousLinearMap.{u2, u2, u1, u2} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (RingHom.id.{u2} 𝕜 (NonAssocRing.toNonAssocSemiring.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u2} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u2} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u2} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2))) (fun (l : ContinuousLinearMap.{u2, u2, u1, u2} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (RingHom.id.{u2} 𝕜 (NonAssocRing.toNonAssocSemiring.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u2} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u2} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u2} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2))) => Exists.{succ u2} 𝕜 (fun (a : 𝕜) => Eq.{succ u1} (Set.{u1} E) B (setOf.{u1} E (fun (x : E) => And (Membership.mem.{u1, u1} E (Set.{u1} E) (Set.instMembershipSet.{u1} E) x A) (LE.le.{u2} 𝕜 (Preorder.toLE.{u2} 𝕜 (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2))) a (FunLike.coe.{max (succ u2) (succ u1), succ u1, succ u2} (ContinuousLinearMap.{u2, u2, u1, u2} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (RingHom.id.{u2} 𝕜 (NonAssocRing.toNonAssocSemiring.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u2} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u2} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u2} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2))) E (fun (a : E) => (fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) a) (ContinuousMapClass.toFunLike.{max u2 u1, u1, u2} (ContinuousLinearMap.{u2, u2, u1, u2} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (RingHom.id.{u2} 𝕜 (NonAssocRing.toNonAssocSemiring.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u2} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u2} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u2} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2))) E 𝕜 _inst_4 _inst_1 (ContinuousSemilinearMapClass.toContinuousMapClass.{max u2 u1, u2, u2, u1, u2} (ContinuousLinearMap.{u2, u2, u1, u2} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (RingHom.id.{u2} 𝕜 (NonAssocRing.toNonAssocSemiring.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u2} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u2} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u2} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2))) 𝕜 𝕜 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (RingHom.id.{u2} 𝕜 (NonAssocRing.toNonAssocSemiring.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u2} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u2} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u2} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2)) (ContinuousLinearMap.continuousSemilinearMapClass.{u2, u2, u1, u2} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (RingHom.id.{u2} 𝕜 (NonAssocRing.toNonAssocSemiring.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u2} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u2} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u2} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2))))) l x))))))) +Case conversion may be inaccurate. Consider using '#align is_exposed.eq_inter_halfspace IsExposed.eq_inter_halfspaceₓ'. -/ /-- For nontrivial `𝕜`, if `B` is an exposed subset of `A`, then `B` is the intersection of `A` with some closed halfspace. The converse is *not* true. It would require that the corresponding open halfspace doesn't intersect `A`. -/ @@ -148,6 +196,12 @@ theorem eq_inter_halfspace [Nontrivial 𝕜] {A B : Set E} (hAB : IsExposed 𝕜 exact hAB.eq_inter_halfspace' hB #align is_exposed.eq_inter_halfspace IsExposed.eq_inter_halfspace +/- warning: is_exposed.inter -> IsExposed.inter is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] [_inst_6 : ContinuousAdd.{u1} 𝕜 _inst_1 (Distrib.toHasAdd.{u1} 𝕜 (Ring.toDistrib.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))] {A : Set.{u2} E} {B : Set.{u2} E} {C : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B) -> (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A C) -> (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A (Inter.inter.{u2} (Set.{u2} E) (Set.hasInter.{u2} E) B C)) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] [_inst_6 : ContinuousAdd.{u2} 𝕜 _inst_1 (Distrib.toAdd.{u2} 𝕜 (NonUnitalNonAssocSemiring.toDistrib.{u2} 𝕜 (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u2} 𝕜 (NonAssocRing.toNonUnitalNonAssocRing.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2))))))] {A : Set.{u1} E} {B : Set.{u1} E} {C : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A B) -> (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A C) -> (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A (Inter.inter.{u1} (Set.{u1} E) (Set.instInterSet.{u1} E) B C)) +Case conversion may be inaccurate. Consider using '#align is_exposed.inter IsExposed.interₓ'. -/ protected theorem inter [ContinuousAdd 𝕜] {A B C : Set E} (hB : IsExposed 𝕜 A B) (hC : IsExposed 𝕜 A C) : IsExposed 𝕜 A (B ∩ C) := by @@ -167,6 +221,12 @@ protected theorem inter [ContinuousAdd 𝕜] {A B C : Set E} (hB : IsExposed (add_le_add_iff_left (l₁ x)).1 (le_trans (add_le_add (hwB.2 x hxA) (hwC.2 y hy)) (hx w hwB.1)) #align is_exposed.inter IsExposed.inter +/- warning: is_exposed.sInter -> IsExposed.interₛ is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} [_inst_6 : ContinuousAdd.{u1} 𝕜 _inst_1 (Distrib.toHasAdd.{u1} 𝕜 (Ring.toDistrib.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))] {F : Finset.{u2} (Set.{u2} E)}, (Finset.Nonempty.{u2} (Set.{u2} E) F) -> (forall (B : Set.{u2} E), (Membership.Mem.{u2, u2} (Set.{u2} E) (Finset.{u2} (Set.{u2} E)) (Finset.hasMem.{u2} (Set.{u2} E)) B F) -> (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B)) -> (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A (Set.interₛ.{u2} E ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} (Set.{u2} E)) (Set.{u2} (Set.{u2} E)) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} (Set.{u2} E)) (Set.{u2} (Set.{u2} E)) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} (Set.{u2} E)) (Set.{u2} (Set.{u2} E)) (Finset.Set.hasCoeT.{u2} (Set.{u2} E)))) F))) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] {A : Set.{u1} E} [_inst_6 : ContinuousAdd.{u2} 𝕜 _inst_1 (Distrib.toAdd.{u2} 𝕜 (NonUnitalNonAssocSemiring.toDistrib.{u2} 𝕜 (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u2} 𝕜 (NonAssocRing.toNonUnitalNonAssocRing.{u2} 𝕜 (Ring.toNonAssocRing.{u2} 𝕜 (OrderedRing.toRing.{u2} 𝕜 _inst_2))))))] {F : Finset.{u1} (Set.{u1} E)}, (Finset.Nonempty.{u1} (Set.{u1} E) F) -> (forall (B : Set.{u1} E), (Membership.mem.{u1, u1} (Set.{u1} E) (Finset.{u1} (Set.{u1} E)) (Finset.instMembershipFinset.{u1} (Set.{u1} E)) B F) -> (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A B)) -> (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A (Set.interₛ.{u1} E (Finset.toSet.{u1} (Set.{u1} E) F))) +Case conversion may be inaccurate. Consider using '#align is_exposed.sInter IsExposed.interₛₓ'. -/ theorem interₛ [ContinuousAdd 𝕜] {F : Finset (Set E)} (hF : F.Nonempty) (hAF : ∀ B ∈ F, IsExposed 𝕜 A B) : IsExposed 𝕜 A (⋂₀ F) := by @@ -185,6 +245,12 @@ theorem interₛ [ContinuousAdd 𝕜] {F : Finset (Set E)} (hF : F.Nonempty) (hF hFnemp fun B hB => hCF B (Finset.mem_insert_of_mem hB)) #align is_exposed.sInter IsExposed.interₛ +/- warning: is_exposed.inter_left -> IsExposed.inter_left is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} {B : Set.{u2} E} {C : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A C) -> (HasSubset.Subset.{u2} (Set.{u2} E) (Set.hasSubset.{u2} E) C B) -> (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 (Inter.inter.{u2} (Set.{u2} E) (Set.hasInter.{u2} E) A B) C) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] {A : Set.{u1} E} {B : Set.{u1} E} {C : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A C) -> (HasSubset.Subset.{u1} (Set.{u1} E) (Set.instHasSubsetSet.{u1} E) C B) -> (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 (Inter.inter.{u1} (Set.{u1} E) (Set.instInterSet.{u1} E) A B) C) +Case conversion may be inaccurate. Consider using '#align is_exposed.inter_left IsExposed.inter_leftₓ'. -/ theorem inter_left (hC : IsExposed 𝕜 A C) (hCB : C ⊆ B) : IsExposed 𝕜 (A ∩ B) C := by rintro ⟨w, hw⟩ @@ -195,12 +261,24 @@ theorem inter_left (hC : IsExposed 𝕜 A C) (hCB : C ⊆ B) : IsExposed 𝕜 (A fun x ⟨⟨hxC, _⟩, hx⟩ => ⟨hxC, fun y hy => (hw.2 y hy).trans (hx w ⟨hC.subset hw, hCB hw⟩)⟩⟩ #align is_exposed.inter_left IsExposed.inter_left +/- warning: is_exposed.inter_right -> IsExposed.inter_right is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} {B : Set.{u2} E} {C : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 B C) -> (HasSubset.Subset.{u2} (Set.{u2} E) (Set.hasSubset.{u2} E) C A) -> (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 (Inter.inter.{u2} (Set.{u2} E) (Set.hasInter.{u2} E) A B) C) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] {A : Set.{u1} E} {B : Set.{u1} E} {C : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 B C) -> (HasSubset.Subset.{u1} (Set.{u1} E) (Set.instHasSubsetSet.{u1} E) C A) -> (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 (Inter.inter.{u1} (Set.{u1} E) (Set.instInterSet.{u1} E) A B) C) +Case conversion may be inaccurate. Consider using '#align is_exposed.inter_right IsExposed.inter_rightₓ'. -/ theorem inter_right (hC : IsExposed 𝕜 B C) (hCA : C ⊆ A) : IsExposed 𝕜 (A ∩ B) C := by rw [inter_comm] exact hC.inter_left hCA #align is_exposed.inter_right IsExposed.inter_right +/- warning: is_exposed.is_closed -> IsExposed.isClosed is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] [_inst_6 : OrderClosedTopology.{u1} 𝕜 _inst_1 (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2)))] {A : Set.{u2} E} {B : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B) -> (IsClosed.{u2} E _inst_4 A) -> (IsClosed.{u2} E _inst_4 B) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] [_inst_6 : OrderClosedTopology.{u2} 𝕜 _inst_1 (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2))] {A : Set.{u1} E} {B : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A B) -> (IsClosed.{u1} E _inst_4 A) -> (IsClosed.{u1} E _inst_4 B) +Case conversion may be inaccurate. Consider using '#align is_exposed.is_closed IsExposed.isClosedₓ'. -/ protected theorem isClosed [OrderClosedTopology 𝕜] {A B : Set E} (hAB : IsExposed 𝕜 A B) (hA : IsClosed A) : IsClosed B := by @@ -210,6 +288,12 @@ protected theorem isClosed [OrderClosedTopology 𝕜] {A B : Set E} (hAB : IsExp exact hA.is_closed_le continuousOn_const l.continuous.continuous_on #align is_exposed.is_closed IsExposed.isClosed +/- warning: is_exposed.is_compact -> IsExposed.isCompact is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] [_inst_6 : OrderClosedTopology.{u1} 𝕜 _inst_1 (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2)))] [_inst_7 : T2Space.{u2} E _inst_4] {A : Set.{u2} E} {B : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B) -> (IsCompact.{u2} E _inst_4 A) -> (IsCompact.{u2} E _inst_4 B) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : OrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) _inst_3] [_inst_6 : OrderClosedTopology.{u2} 𝕜 _inst_1 (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2))] [_inst_7 : T2Space.{u1} E _inst_4] {A : Set.{u1} E} {B : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (OrderedSemiring.toSemiring.{u2} 𝕜 (OrderedRing.toOrderedSemiring.{u2} 𝕜 _inst_2)) (PartialOrder.toPreorder.{u2} 𝕜 (OrderedRing.toPartialOrder.{u2} 𝕜 _inst_2)) _inst_3 _inst_4 _inst_5 A B) -> (IsCompact.{u1} E _inst_4 A) -> (IsCompact.{u1} E _inst_4 B) +Case conversion may be inaccurate. Consider using '#align is_exposed.is_compact IsExposed.isCompactₓ'. -/ protected theorem isCompact [OrderClosedTopology 𝕜] [T2Space E] {A B : Set E} (hAB : IsExposed 𝕜 A B) (hA : IsCompact A) : IsCompact B := isCompact_of_isClosed_subset hA (hAB.IsClosed hA.IsClosed) hAB.Subset @@ -219,27 +303,40 @@ end IsExposed variable (𝕜) +#print Set.exposedPoints /- /-- A point is exposed with respect to `A` iff there exists an hyperplane whose intersection with `A` is exactly that point. -/ def Set.exposedPoints (A : Set E) : Set E := { x ∈ A | ∃ l : E →L[𝕜] 𝕜, ∀ y ∈ A, l y ≤ l x ∧ (l x ≤ l y → y = x) } #align set.exposed_points Set.exposedPoints +-/ variable {𝕜} +/- warning: exposed_point_def -> exposed_point_def is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} {x : E}, Iff (Membership.Mem.{u2, u2} E (Set.{u2} E) (Set.hasMem.{u2} E) x (Set.exposedPoints.{u1, u2} 𝕜 E _inst_1 _inst_2 _inst_3 _inst_4 _inst_5 A)) (And (Membership.Mem.{u2, u2} E (Set.{u2} E) (Set.hasMem.{u2} E) x A) (Exists.{max (succ u2) (succ u1)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) (fun (l : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) => forall (y : E), (Membership.Mem.{u2, u2} E (Set.{u2} E) (Set.hasMem.{u2} E) y A) -> (And (LE.le.{u1} 𝕜 (Preorder.toLE.{u1} 𝕜 (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2)))) (coeFn.{max (succ u2) (succ u1), max (succ u2) (succ u1)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) (fun (_x : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) => E -> 𝕜) (ContinuousLinearMap.toFun.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) l y) (coeFn.{max (succ u2) (succ u1), max (succ u2) (succ u1)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) (fun (_x : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) => E -> 𝕜) (ContinuousLinearMap.toFun.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) l x)) ((LE.le.{u1} 𝕜 (Preorder.toLE.{u1} 𝕜 (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2)))) (coeFn.{max (succ u2) (succ u1), max (succ u2) (succ u1)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) (fun (_x : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) => E -> 𝕜) (ContinuousLinearMap.toFun.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) l x) (coeFn.{max (succ u2) (succ u1), max (succ u2) (succ u1)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) (fun (_x : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) => E -> 𝕜) (ContinuousLinearMap.toFun.{u1, u1, u2, u1} 𝕜 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (Semiring.toNonAssocSemiring.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (AddCommGroup.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toAddCommGroup.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (Semiring.toModule.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) l y)) -> (Eq.{succ u2} E y x)))))) +but is expected to have type + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : OrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) _inst_3] {A : Set.{u2} E} {x : E}, Iff (Membership.mem.{u2, u2} E (Set.{u2} E) (Set.instMembershipSet.{u2} E) x (Set.exposedPoints.{u1, u2} 𝕜 E _inst_1 _inst_2 _inst_3 _inst_4 _inst_5 A)) (And (Membership.mem.{u2, u2} E (Set.{u2} E) (Set.instMembershipSet.{u2} E) x A) (Exists.{max (succ u1) (succ u2)} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) (fun (l : ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) => forall (y : E), (Membership.mem.{u2, u2} E (Set.{u2} E) (Set.instMembershipSet.{u2} E) y A) -> (And (LE.le.{u1} ((fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) y) (Preorder.toLE.{u1} ((fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) y) (PartialOrder.toPreorder.{u1} ((fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) y) (OrderedRing.toPartialOrder.{u1} ((fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) y) _inst_2))) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E (fun (_x : E) => (fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) _x) (ContinuousMapClass.toFunLike.{max u1 u2, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E 𝕜 _inst_4 _inst_1 (ContinuousSemilinearMapClass.toContinuousMapClass.{max u1 u2, u1, u1, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (ContinuousLinearMap.continuousSemilinearMapClass.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))))) l y) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E (fun (_x : E) => (fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) _x) (ContinuousMapClass.toFunLike.{max u1 u2, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E 𝕜 _inst_4 _inst_1 (ContinuousSemilinearMapClass.toContinuousMapClass.{max u1 u2, u1, u1, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (ContinuousLinearMap.continuousSemilinearMapClass.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))))) l x)) ((LE.le.{u1} ((fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) x) (Preorder.toLE.{u1} ((fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) x) (PartialOrder.toPreorder.{u1} ((fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) x) (OrderedRing.toPartialOrder.{u1} ((fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) x) _inst_2))) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E (fun (_x : E) => (fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) _x) (ContinuousMapClass.toFunLike.{max u1 u2, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E 𝕜 _inst_4 _inst_1 (ContinuousSemilinearMapClass.toContinuousMapClass.{max u1 u2, u1, u1, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (ContinuousLinearMap.continuousSemilinearMapClass.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))))) l x) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E (fun (_x : E) => (fun (x._@.Mathlib.Topology.ContinuousFunction.Basic._hyg.699 : E) => 𝕜) _x) (ContinuousMapClass.toFunLike.{max u1 u2, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) E 𝕜 _inst_4 _inst_1 (ContinuousSemilinearMapClass.toContinuousMapClass.{max u1 u2, u1, u1, u2, u1} (ContinuousLinearMap.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))) 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)) (ContinuousLinearMap.continuousSemilinearMapClass.{u1, u1, u2, u1} 𝕜 𝕜 (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (OrderedSemiring.toSemiring.{u1} 𝕜 (OrderedRing.toOrderedSemiring.{u1} 𝕜 _inst_2)) (RingHom.id.{u1} 𝕜 (NonAssocRing.toNonAssocSemiring.{u1} 𝕜 (Ring.toNonAssocRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2)))) E _inst_4 _inst_3 𝕜 _inst_1 (OrderedCancelAddCommMonoid.toAddCommMonoid.{u1} 𝕜 (OrderedAddCommGroup.toOrderedCancelAddCommMonoid.{u1} 𝕜 (OrderedRing.toOrderedAddCommGroup.{u1} 𝕜 _inst_2))) _inst_5 (instModuleToSemiringToAddCommMonoidToNonUnitalNonAssocSemiringToNonUnitalNonAssocRingToNonUnitalRing.{u1} 𝕜 (OrderedRing.toRing.{u1} 𝕜 _inst_2))))) l y)) -> (Eq.{succ u2} E y x)))))) +Case conversion may be inaccurate. Consider using '#align exposed_point_def exposed_point_defₓ'. -/ theorem exposed_point_def : x ∈ A.exposedPoints 𝕜 ↔ x ∈ A ∧ ∃ l : E →L[𝕜] 𝕜, ∀ y ∈ A, l y ≤ l x ∧ (l x ≤ l y → y = x) := Iff.rfl #align exposed_point_def exposed_point_def +#print exposedPoints_subset /- theorem exposedPoints_subset : A.exposedPoints 𝕜 ⊆ A := fun x hx => hx.1 #align exposed_points_subset exposedPoints_subset +-/ +#print exposedPoints_empty /- @[simp] theorem exposedPoints_empty : (∅ : Set E).exposedPoints 𝕜 = ∅ := subset_empty_iff.1 exposedPoints_subset #align exposed_points_empty exposedPoints_empty +-/ +#print mem_exposedPoints_iff_exposed_singleton /- /-- Exposed points exactly correspond to exposed singletons. -/ theorem mem_exposedPoints_iff_exposed_singleton : x ∈ A.exposedPoints 𝕜 ↔ IsExposed 𝕜 A {x} := by @@ -255,6 +352,7 @@ theorem mem_exposedPoints_iff_exposed_singleton : x ∈ A.exposedPoints 𝕜 ↔ ⟨hl.1.1, l, fun y hy => ⟨hl.1.2 y hy, fun hxy => hl.2 y ⟨hy, fun z hz => (hl.1.2 z hz).trans hxy⟩⟩⟩ #align mem_exposed_points_iff_exposed_singleton mem_exposedPoints_iff_exposed_singleton +-/ end OrderedRing @@ -265,6 +363,12 @@ variable {𝕜 : Type _} {E : Type _} [TopologicalSpace 𝕜] [LinearOrderedRing namespace IsExposed +/- warning: is_exposed.convex -> IsExposed.convex is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : LinearOrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) _inst_3] {A : Set.{u2} E} {B : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (StrictOrderedRing.toOrderedAddCommGroup.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2)))) _inst_3 _inst_4 _inst_5 A B) -> (Convex.{u1, u2} 𝕜 E (StrictOrderedSemiring.toOrderedSemiring.{u1} 𝕜 (StrictOrderedRing.toStrictOrderedSemiring.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) _inst_3 (SMulZeroClass.toHasSmul.{u1, u2} 𝕜 E (AddZeroClass.toHasZero.{u2} E (AddMonoid.toAddZeroClass.{u2} E (AddCommMonoid.toAddMonoid.{u2} E _inst_3))) (SMulWithZero.toSmulZeroClass.{u1, u2} 𝕜 E (MulZeroClass.toHasZero.{u1} 𝕜 (MulZeroOneClass.toMulZeroClass.{u1} 𝕜 (MonoidWithZero.toMulZeroOneClass.{u1} 𝕜 (Semiring.toMonoidWithZero.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))))))) (AddZeroClass.toHasZero.{u2} E (AddMonoid.toAddZeroClass.{u2} E (AddCommMonoid.toAddMonoid.{u2} E _inst_3))) (MulActionWithZero.toSMulWithZero.{u1, u2} 𝕜 E (Semiring.toMonoidWithZero.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2)))) (AddZeroClass.toHasZero.{u2} E (AddMonoid.toAddZeroClass.{u2} E (AddCommMonoid.toAddMonoid.{u2} E _inst_3))) (Module.toMulActionWithZero.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) _inst_3 _inst_5)))) A) -> (Convex.{u1, u2} 𝕜 E (StrictOrderedSemiring.toOrderedSemiring.{u1} 𝕜 (StrictOrderedRing.toStrictOrderedSemiring.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) _inst_3 (SMulZeroClass.toHasSmul.{u1, u2} 𝕜 E (AddZeroClass.toHasZero.{u2} E (AddMonoid.toAddZeroClass.{u2} E (AddCommMonoid.toAddMonoid.{u2} E _inst_3))) (SMulWithZero.toSmulZeroClass.{u1, u2} 𝕜 E (MulZeroClass.toHasZero.{u1} 𝕜 (MulZeroOneClass.toMulZeroClass.{u1} 𝕜 (MonoidWithZero.toMulZeroOneClass.{u1} 𝕜 (Semiring.toMonoidWithZero.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))))))) (AddZeroClass.toHasZero.{u2} E (AddMonoid.toAddZeroClass.{u2} E (AddCommMonoid.toAddMonoid.{u2} E _inst_3))) (MulActionWithZero.toSMulWithZero.{u1, u2} 𝕜 E (Semiring.toMonoidWithZero.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2)))) (AddZeroClass.toHasZero.{u2} E (AddMonoid.toAddZeroClass.{u2} E (AddCommMonoid.toAddMonoid.{u2} E _inst_3))) (Module.toMulActionWithZero.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) _inst_3 _inst_5)))) B) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : LinearOrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) _inst_3] {A : Set.{u1} E} {B : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) (PartialOrder.toPreorder.{u2} 𝕜 (StrictOrderedRing.toPartialOrder.{u2} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u2} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B) -> (Convex.{u2, u1} 𝕜 E (StrictOrderedSemiring.toOrderedSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) _inst_3 (SMulZeroClass.toSMul.{u2, u1} 𝕜 E (AddMonoid.toZero.{u1} E (AddCommMonoid.toAddMonoid.{u1} E _inst_3)) (SMulWithZero.toSMulZeroClass.{u2, u1} 𝕜 E (MonoidWithZero.toZero.{u2} 𝕜 (Semiring.toMonoidWithZero.{u2} 𝕜 (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))))) (AddMonoid.toZero.{u1} E (AddCommMonoid.toAddMonoid.{u1} E _inst_3)) (MulActionWithZero.toSMulWithZero.{u2, u1} 𝕜 E (Semiring.toMonoidWithZero.{u2} 𝕜 (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2)))) (AddMonoid.toZero.{u1} E (AddCommMonoid.toAddMonoid.{u1} E _inst_3)) (Module.toMulActionWithZero.{u2, u1} 𝕜 E (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) _inst_3 _inst_5)))) A) -> (Convex.{u2, u1} 𝕜 E (StrictOrderedSemiring.toOrderedSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) _inst_3 (SMulZeroClass.toSMul.{u2, u1} 𝕜 E (AddMonoid.toZero.{u1} E (AddCommMonoid.toAddMonoid.{u1} E _inst_3)) (SMulWithZero.toSMulZeroClass.{u2, u1} 𝕜 E (MonoidWithZero.toZero.{u2} 𝕜 (Semiring.toMonoidWithZero.{u2} 𝕜 (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))))) (AddMonoid.toZero.{u1} E (AddCommMonoid.toAddMonoid.{u1} E _inst_3)) (MulActionWithZero.toSMulWithZero.{u2, u1} 𝕜 E (Semiring.toMonoidWithZero.{u2} 𝕜 (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2)))) (AddMonoid.toZero.{u1} E (AddCommMonoid.toAddMonoid.{u1} E _inst_3)) (Module.toMulActionWithZero.{u2, u1} 𝕜 E (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) _inst_3 _inst_5)))) B) +Case conversion may be inaccurate. Consider using '#align is_exposed.convex IsExposed.convexₓ'. -/ protected theorem convex (hAB : IsExposed 𝕜 A B) (hA : Convex 𝕜 A) : Convex 𝕜 B := by obtain rfl | hB := B.eq_empty_or_nonempty @@ -276,6 +380,12 @@ protected theorem convex (hAB : IsExposed 𝕜 A B) (hA : Convex 𝕜 A) : Conve ⟨mem_univ _, hx₂.2 y hy⟩ ha hb hab).2⟩ #align is_exposed.convex IsExposed.convex +/- warning: is_exposed.is_extreme -> IsExposed.isExtreme is a dubious translation: +lean 3 declaration is + forall {𝕜 : Type.{u1}} {E : Type.{u2}} [_inst_1 : TopologicalSpace.{u1} 𝕜] [_inst_2 : LinearOrderedRing.{u1} 𝕜] [_inst_3 : AddCommMonoid.{u2} E] [_inst_4 : TopologicalSpace.{u2} E] [_inst_5 : Module.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) _inst_3] {A : Set.{u2} E} {B : Set.{u2} E}, (IsExposed.{u1, u2} 𝕜 E _inst_1 (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) (PartialOrder.toPreorder.{u1} 𝕜 (OrderedAddCommGroup.toPartialOrder.{u1} 𝕜 (StrictOrderedRing.toOrderedAddCommGroup.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2)))) _inst_3 _inst_4 _inst_5 A B) -> (IsExtreme.{u1, u2} 𝕜 E (StrictOrderedSemiring.toOrderedSemiring.{u1} 𝕜 (StrictOrderedRing.toStrictOrderedSemiring.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) _inst_3 (SMulZeroClass.toHasSmul.{u1, u2} 𝕜 E (AddZeroClass.toHasZero.{u2} E (AddMonoid.toAddZeroClass.{u2} E (AddCommMonoid.toAddMonoid.{u2} E _inst_3))) (SMulWithZero.toSmulZeroClass.{u1, u2} 𝕜 E (MulZeroClass.toHasZero.{u1} 𝕜 (MulZeroOneClass.toMulZeroClass.{u1} 𝕜 (MonoidWithZero.toMulZeroOneClass.{u1} 𝕜 (Semiring.toMonoidWithZero.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))))))) (AddZeroClass.toHasZero.{u2} E (AddMonoid.toAddZeroClass.{u2} E (AddCommMonoid.toAddMonoid.{u2} E _inst_3))) (MulActionWithZero.toSMulWithZero.{u1, u2} 𝕜 E (Semiring.toMonoidWithZero.{u1} 𝕜 (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2)))) (AddZeroClass.toHasZero.{u2} E (AddMonoid.toAddZeroClass.{u2} E (AddCommMonoid.toAddMonoid.{u2} E _inst_3))) (Module.toMulActionWithZero.{u1, u2} 𝕜 E (Ring.toSemiring.{u1} 𝕜 (StrictOrderedRing.toRing.{u1} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u1} 𝕜 _inst_2))) _inst_3 _inst_5)))) A B) +but is expected to have type + forall {𝕜 : Type.{u2}} {E : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} 𝕜] [_inst_2 : LinearOrderedRing.{u2} 𝕜] [_inst_3 : AddCommMonoid.{u1} E] [_inst_4 : TopologicalSpace.{u1} E] [_inst_5 : Module.{u2, u1} 𝕜 E (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) _inst_3] {A : Set.{u1} E} {B : Set.{u1} E}, (IsExposed.{u2, u1} 𝕜 E _inst_1 (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) (PartialOrder.toPreorder.{u2} 𝕜 (StrictOrderedRing.toPartialOrder.{u2} 𝕜 (LinearOrderedRing.toStrictOrderedRing.{u2} 𝕜 _inst_2))) _inst_3 _inst_4 _inst_5 A B) -> (IsExtreme.{u2, u1} 𝕜 E (StrictOrderedSemiring.toOrderedSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) _inst_3 (SMulZeroClass.toSMul.{u2, u1} 𝕜 E (AddMonoid.toZero.{u1} E (AddCommMonoid.toAddMonoid.{u1} E _inst_3)) (SMulWithZero.toSMulZeroClass.{u2, u1} 𝕜 E (MonoidWithZero.toZero.{u2} 𝕜 (Semiring.toMonoidWithZero.{u2} 𝕜 (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))))) (AddMonoid.toZero.{u1} E (AddCommMonoid.toAddMonoid.{u1} E _inst_3)) (MulActionWithZero.toSMulWithZero.{u2, u1} 𝕜 E (Semiring.toMonoidWithZero.{u2} 𝕜 (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2)))) (AddMonoid.toZero.{u1} E (AddCommMonoid.toAddMonoid.{u1} E _inst_3)) (Module.toMulActionWithZero.{u2, u1} 𝕜 E (StrictOrderedSemiring.toSemiring.{u2} 𝕜 (LinearOrderedSemiring.toStrictOrderedSemiring.{u2} 𝕜 (LinearOrderedRing.toLinearOrderedSemiring.{u2} 𝕜 _inst_2))) _inst_3 _inst_5)))) A B) +Case conversion may be inaccurate. Consider using '#align is_exposed.is_extreme IsExposed.isExtremeₓ'. -/ protected theorem isExtreme (hAB : IsExposed 𝕜 A B) : IsExtreme 𝕜 A B := by refine' ⟨hAB.subset, fun x₁ hx₁A x₂ hx₂A x hxB hx => _⟩ @@ -293,9 +403,11 @@ protected theorem isExtreme (hAB : IsExposed 𝕜 A B) : IsExtreme 𝕜 A B := end IsExposed +#print exposedPoints_subset_extremePoints /- theorem exposedPoints_subset_extremePoints : A.exposedPoints 𝕜 ⊆ A.extremePoints 𝕜 := fun x hx => mem_extremePoints_iff_extreme_singleton.2 (mem_exposedPoints_iff_exposed_singleton.1 hx).IsExtreme #align exposed_points_subset_extreme_points exposedPoints_subset_extremePoints +-/ end LinearOrderedRing diff --git a/Mathbin/Analysis/Normed/Ring/Seminorm.lean b/Mathbin/Analysis/Normed/Ring/Seminorm.lean index 09645d4970..223069979f 100644 --- a/Mathbin/Analysis/Normed/Ring/Seminorm.lean +++ b/Mathbin/Analysis/Normed/Ring/Seminorm.lean @@ -45,28 +45,36 @@ open NNReal variable {F R S : Type _} (x y : R) (r : ℝ) +#print RingSeminorm /- /-- A seminorm on a ring `R` is a function `f : R → ℝ` that preserves zero, takes nonnegative values, is subadditive and submultiplicative and such that `f (-x) = f x` for all `x ∈ R`. -/ structure RingSeminorm (R : Type _) [NonUnitalNonAssocRing R] extends AddGroupSeminorm R where mul_le' : ∀ x y : R, to_fun (x * y) ≤ to_fun x * to_fun y #align ring_seminorm RingSeminorm +-/ +#print RingNorm /- /-- A function `f : R → ℝ` is a norm on a (nonunital) ring if it is a seminorm and `f x = 0` implies `x = 0`. -/ structure RingNorm (R : Type _) [NonUnitalNonAssocRing R] extends RingSeminorm R, AddGroupNorm R #align ring_norm RingNorm +-/ +#print MulRingSeminorm /- /-- A multiplicative seminorm on a ring `R` is a function `f : R → ℝ` that preserves zero and multiplication, takes nonnegative values, is subadditive and such that `f (-x) = f x` for all `x`. -/ structure MulRingSeminorm (R : Type _) [NonAssocRing R] extends AddGroupSeminorm R, MonoidWithZeroHom R ℝ #align mul_ring_seminorm MulRingSeminorm +-/ +#print MulRingNorm /- /-- A multiplicative norm on a ring `R` is a multiplicative ring seminorm such that `f x = 0` implies `x = 0`. -/ structure MulRingNorm (R : Type _) [NonAssocRing R] extends MulRingSeminorm R, AddGroupNorm R #align mul_ring_norm MulRingNorm +-/ attribute [nolint doc_blame] RingSeminorm.toAddGroupSeminorm RingNorm.toAddGroupNorm RingNorm.toRingSeminorm MulRingSeminorm.toAddGroupSeminorm MulRingSeminorm.toMonoidWithZeroHom MulRingNorm.toAddGroupNorm MulRingNorm.toMulRingSeminorm @@ -77,6 +85,7 @@ section NonUnitalRing variable [NonUnitalRing R] +#print RingSeminorm.ringSeminormClass /- instance ringSeminormClass : RingSeminormClass (RingSeminorm R) R ℝ where coe f := f.toFun @@ -86,16 +95,29 @@ instance ringSeminormClass : RingSeminormClass (RingSeminorm R) R ℝ map_mul_le_mul f := f.mul_le' map_neg_eq_map f := f.neg' #align ring_seminorm.ring_seminorm_class RingSeminorm.ringSeminormClass +-/ /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ instance : CoeFun (RingSeminorm R) fun _ => R → ℝ := FunLike.hasCoeToFun +/- warning: ring_seminorm.to_fun_eq_coe -> RingSeminorm.toFun_eq_coe is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] (p : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)), Eq.{succ u1} (R -> Real) (RingSeminorm.toFun.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) p) (coeFn.{succ u1, succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (fun (_x : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) => R -> Real) (RingSeminorm.hasCoeToFun.{u1} R _inst_1) p) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] (p : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)), Eq.{succ u1} (R -> Real) (FunLike.coe.{succ u1, succ u1, 1} (AddGroupSeminorm.{u1} R (AddCommGroup.toAddGroup.{u1} R (NonUnitalNonAssocRing.toAddCommGroup.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))) R (fun (a._@.Mathlib.Analysis.Normed.Group.Seminorm._hyg.956 : R) => Real) (SubadditiveHomClass.toFunLike.{u1, u1, 0} (AddGroupSeminorm.{u1} R (AddCommGroup.toAddGroup.{u1} R (NonUnitalNonAssocRing.toAddCommGroup.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))) R Real (AddZeroClass.toAdd.{u1} R (AddMonoid.toAddZeroClass.{u1} R (SubNegMonoid.toAddMonoid.{u1} R (AddGroup.toSubNegMonoid.{u1} R (AddCommGroup.toAddGroup.{u1} R (NonUnitalNonAssocRing.toAddCommGroup.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1))))))) (AddZeroClass.toAdd.{0} Real (AddMonoid.toAddZeroClass.{0} Real (AddCommMonoid.toAddMonoid.{0} Real (OrderedAddCommMonoid.toAddCommMonoid.{0} Real Real.orderedAddCommMonoid)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedAddCommMonoid.toPartialOrder.{0} Real Real.orderedAddCommMonoid))) (AddGroupSeminormClass.toSubadditiveHomClass.{u1, u1, 0} (AddGroupSeminorm.{u1} R (AddCommGroup.toAddGroup.{u1} R (NonUnitalNonAssocRing.toAddCommGroup.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))) R Real (AddCommGroup.toAddGroup.{u1} R (NonUnitalNonAssocRing.toAddCommGroup.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1))) Real.orderedAddCommMonoid (AddGroupSeminorm.addGroupSeminormClass.{u1} R (AddCommGroup.toAddGroup.{u1} R (NonUnitalNonAssocRing.toAddCommGroup.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))))) (RingSeminorm.toAddGroupSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) p)) (FunLike.coe.{succ u1, succ u1, 1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (RingSeminorm.ringSeminormClass.{u1} R _inst_1))) p) +Case conversion may be inaccurate. Consider using '#align ring_seminorm.to_fun_eq_coe RingSeminorm.toFun_eq_coeₓ'. -/ @[simp] theorem toFun_eq_coe (p : RingSeminorm R) : p.toFun = p := rfl #align ring_seminorm.to_fun_eq_coe RingSeminorm.toFun_eq_coe +/- warning: ring_seminorm.ext -> RingSeminorm.ext is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] {p : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)} {q : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)}, (forall (x : R), Eq.{1} Real (coeFn.{succ u1, succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (fun (_x : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) => R -> Real) (RingSeminorm.hasCoeToFun.{u1} R _inst_1) p x) (coeFn.{succ u1, succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (fun (_x : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) => R -> Real) (RingSeminorm.hasCoeToFun.{u1} R _inst_1) q x)) -> (Eq.{succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) p q) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] {p : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)} {q : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)}, (forall (x : R), Eq.{1} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) (FunLike.coe.{succ u1, succ u1, 1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (RingSeminorm.ringSeminormClass.{u1} R _inst_1))) p x) (FunLike.coe.{succ u1, succ u1, 1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (RingSeminorm.ringSeminormClass.{u1} R _inst_1))) q x)) -> (Eq.{succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) p q) +Case conversion may be inaccurate. Consider using '#align ring_seminorm.ext RingSeminorm.extₓ'. -/ @[ext] theorem ext {p q : RingSeminorm R} : (∀ x, p x = q x) → p = q := FunLike.ext p q @@ -104,10 +126,22 @@ theorem ext {p q : RingSeminorm R} : (∀ x, p x = q x) → p = q := instance : Zero (RingSeminorm R) := ⟨{ AddGroupSeminorm.hasZero.zero with mul_le' := fun _ _ => (MulZeroClass.zero_mul _).ge }⟩ +/- warning: ring_seminorm.eq_zero_iff -> RingSeminorm.eq_zero_iff is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] {p : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)}, Iff (Eq.{succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) p (OfNat.ofNat.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 0 (OfNat.mk.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 0 (Zero.zero.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (RingSeminorm.hasZero.{u1} R _inst_1))))) (forall (x : R), Eq.{1} Real (coeFn.{succ u1, succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (fun (_x : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) => R -> Real) (RingSeminorm.hasCoeToFun.{u1} R _inst_1) p x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] {p : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)}, Iff (Eq.{succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) p (OfNat.ofNat.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 0 (Zero.toOfNat0.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (RingSeminorm.instZeroRingSeminormToNonUnitalNonAssocRing.{u1} R _inst_1)))) (forall (x : R), Eq.{1} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) (FunLike.coe.{succ u1, succ u1, 1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (RingSeminorm.ringSeminormClass.{u1} R _inst_1))) p x) (OfNat.ofNat.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) 0 (Zero.toOfNat0.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) Real.instZeroReal))) +Case conversion may be inaccurate. Consider using '#align ring_seminorm.eq_zero_iff RingSeminorm.eq_zero_iffₓ'. -/ theorem eq_zero_iff {p : RingSeminorm R} : p = 0 ↔ ∀ x, p x = 0 := FunLike.ext_iff #align ring_seminorm.eq_zero_iff RingSeminorm.eq_zero_iff +/- warning: ring_seminorm.ne_zero_iff -> RingSeminorm.ne_zero_iff is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] {p : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)}, Iff (Ne.{succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) p (OfNat.ofNat.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 0 (OfNat.mk.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 0 (Zero.zero.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (RingSeminorm.hasZero.{u1} R _inst_1))))) (Exists.{succ u1} R (fun (x : R) => Ne.{1} Real (coeFn.{succ u1, succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (fun (_x : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) => R -> Real) (RingSeminorm.hasCoeToFun.{u1} R _inst_1) p x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] {p : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)}, Iff (Ne.{succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) p (OfNat.ofNat.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 0 (Zero.toOfNat0.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (RingSeminorm.instZeroRingSeminormToNonUnitalNonAssocRing.{u1} R _inst_1)))) (Exists.{succ u1} R (fun (x : R) => Ne.{1} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) (FunLike.coe.{succ u1, succ u1, 1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (RingSeminorm.ringSeminormClass.{u1} R _inst_1))) p x) (OfNat.ofNat.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) 0 (Zero.toOfNat0.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) Real.instZeroReal)))) +Case conversion may be inaccurate. Consider using '#align ring_seminorm.ne_zero_iff RingSeminorm.ne_zero_iffₓ'. -/ theorem ne_zero_iff {p : RingSeminorm R} : p ≠ 0 ↔ ∃ x, p x ≠ 0 := by simp [eq_zero_iff] #align ring_seminorm.ne_zero_iff RingSeminorm.ne_zero_iff @@ -128,6 +162,12 @@ instance [DecidableEq R] : One (RingSeminorm R) := · change ite _ _ _ ≤ ite _ _ _ * ite _ _ _ simp only [if_false, h, left_ne_zero_of_mul h, right_ne_zero_of_mul h, mul_one] }⟩ +/- warning: ring_seminorm.apply_one -> RingSeminorm.apply_one is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] [_inst_2 : DecidableEq.{succ u1} R] (x : R), Eq.{1} Real (coeFn.{succ u1, succ u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (fun (_x : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) => R -> Real) (RingSeminorm.hasCoeToFun.{u1} R _inst_1) (OfNat.ofNat.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 1 (OfNat.mk.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 1 (One.one.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (RingSeminorm.hasOne.{u1} R _inst_1 (fun (a : R) (b : R) => _inst_2 a b))))) x) (ite.{1} Real (Eq.{succ u1} R x (OfNat.ofNat.{u1} R 0 (OfNat.mk.{u1} R 0 (Zero.zero.{u1} R (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))))))) (_inst_2 x (OfNat.ofNat.{u1} R 0 (OfNat.mk.{u1} R 0 (Zero.zero.{u1} R (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))))))) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] [_inst_2 : DecidableEq.{succ u1} R] (x : R), Eq.{1} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) (FunLike.coe.{succ u1, succ u1, 1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) R Real (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (RingSeminorm.ringSeminormClass.{u1} R _inst_1))) (OfNat.ofNat.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 1 (One.toOfNat1.{u1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (RingSeminorm.instOneRingSeminormToNonUnitalNonAssocRing.{u1} R _inst_1 (fun (a : R) (b : R) => _inst_2 a b)))) x) (ite.{1} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) (Eq.{succ u1} R x (OfNat.ofNat.{u1} R 0 (Zero.toOfNat0.{u1} R (SemigroupWithZero.toZero.{u1} R (NonUnitalSemiring.toSemigroupWithZero.{u1} R (NonUnitalRing.toNonUnitalSemiring.{u1} R _inst_1)))))) (_inst_2 x (OfNat.ofNat.{u1} R 0 (Zero.toOfNat0.{u1} R (SemigroupWithZero.toZero.{u1} R (NonUnitalSemiring.toSemigroupWithZero.{u1} R (NonUnitalRing.toNonUnitalSemiring.{u1} R _inst_1)))))) (OfNat.ofNat.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) 0 (Zero.toOfNat0.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) Real.instZeroReal)) (OfNat.ofNat.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) 1 (One.toOfNat1.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) x) Real.instOneReal))) +Case conversion may be inaccurate. Consider using '#align ring_seminorm.apply_one RingSeminorm.apply_oneₓ'. -/ @[simp] theorem apply_one [DecidableEq R] (x : R) : (1 : RingSeminorm R) x = if x = 0 then 0 else 1 := rfl @@ -139,6 +179,12 @@ section Ring variable [Ring R] (p : RingSeminorm R) +/- warning: ring_seminorm.seminorm_one_eq_one_iff_ne_zero -> RingSeminorm.seminorm_one_eq_one_iff_ne_zero is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : Ring.{u1} R] (p : RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))), (LE.le.{0} Real Real.hasLe (coeFn.{succ u1, succ u1} (RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))) (fun (_x : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) => R -> Real) (RingSeminorm.hasCoeToFun.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1)) p (OfNat.ofNat.{u1} R 1 (OfNat.mk.{u1} R 1 (One.one.{u1} R (AddMonoidWithOne.toOne.{u1} R (AddGroupWithOne.toAddMonoidWithOne.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (Ring.toAddCommGroupWithOne.{u1} R _inst_1)))))))) (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Iff (Eq.{1} Real (coeFn.{succ u1, succ u1} (RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))) (fun (_x : RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) => R -> Real) (RingSeminorm.hasCoeToFun.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1)) p (OfNat.ofNat.{u1} R 1 (OfNat.mk.{u1} R 1 (One.one.{u1} R (AddMonoidWithOne.toOne.{u1} R (AddGroupWithOne.toAddMonoidWithOne.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (Ring.toAddCommGroupWithOne.{u1} R _inst_1)))))))) (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) (Ne.{succ u1} (RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))) p (OfNat.ofNat.{u1} (RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))) 0 (OfNat.mk.{u1} (RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))) 0 (Zero.zero.{u1} (RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))) (RingSeminorm.hasZero.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))))))) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : Ring.{u1} R] (p : RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))), (LE.le.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) (OfNat.ofNat.{u1} R 1 (One.toOfNat1.{u1} R (NonAssocRing.toOne.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))))) Real.instLEReal (FunLike.coe.{succ u1, succ u1, 1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) R Real (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1)) Real.orderedSemiring (RingSeminorm.ringSeminormClass.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1)))) p (OfNat.ofNat.{u1} R 1 (One.toOfNat1.{u1} R (NonAssocRing.toOne.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))))) (OfNat.ofNat.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) (OfNat.ofNat.{u1} R 1 (One.toOfNat1.{u1} R (NonAssocRing.toOne.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))))) 1 (One.toOfNat1.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) (OfNat.ofNat.{u1} R 1 (One.toOfNat1.{u1} R (NonAssocRing.toOne.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))))) Real.instOneReal))) -> (Iff (Eq.{1} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) (OfNat.ofNat.{u1} R 1 (One.toOfNat1.{u1} R (NonAssocRing.toOne.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))))) (FunLike.coe.{succ u1, succ u1, 1} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (RingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1))) R Real (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1)) Real.orderedSemiring (RingSeminorm.ringSeminormClass.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1)))) p (OfNat.ofNat.{u1} R 1 (One.toOfNat1.{u1} R (NonAssocRing.toOne.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))))) (OfNat.ofNat.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) (OfNat.ofNat.{u1} R 1 (One.toOfNat1.{u1} R (NonAssocRing.toOne.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))))) 1 (One.toOfNat1.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.319 : R) => Real) (OfNat.ofNat.{u1} R 1 (One.toOfNat1.{u1} R (NonAssocRing.toOne.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))))) Real.instOneReal))) (Ne.{succ u1} (RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))) p (OfNat.ofNat.{u1} (RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))) 0 (Zero.toOfNat0.{u1} (RingSeminorm.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R (Ring.toNonAssocRing.{u1} R _inst_1))) (RingSeminorm.instZeroRingSeminormToNonUnitalNonAssocRing.{u1} R (Ring.toNonUnitalRing.{u1} R _inst_1)))))) +Case conversion may be inaccurate. Consider using '#align ring_seminorm.seminorm_one_eq_one_iff_ne_zero RingSeminorm.seminorm_one_eq_one_iff_ne_zeroₓ'. -/ theorem seminorm_one_eq_one_iff_ne_zero (hp : p 1 ≤ 1) : p 1 = 1 ↔ p ≠ 0 := by refine' @@ -159,17 +205,20 @@ end Ring end RingSeminorm +#print normRingSeminorm /- /-- The norm of a `non_unital_semi_normed_ring` as a `ring_seminorm`. -/ def normRingSeminorm (R : Type _) [NonUnitalSeminormedRing R] : RingSeminorm R := { normAddGroupSeminorm R with toFun := norm mul_le' := norm_mul_le } #align norm_ring_seminorm normRingSeminorm +-/ namespace RingNorm variable [NonUnitalRing R] +#print RingNorm.ringNormClass /- instance ringNormClass : RingNormClass (RingNorm R) R ℝ where coe f := f.toFun @@ -180,16 +229,30 @@ instance ringNormClass : RingNormClass (RingNorm R) R ℝ map_neg_eq_map f := f.neg' eq_zero_of_map_eq_zero f := f.eq_zero_of_map_eq_zero' #align ring_norm.ring_norm_class RingNorm.ringNormClass +-/ /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ instance : CoeFun (RingNorm R) fun _ => R → ℝ := ⟨fun p => p.toFun⟩ +/- warning: ring_norm.to_fun_eq_coe clashes with [anonymous] -> [anonymous] +warning: ring_norm.to_fun_eq_coe -> [anonymous] is a dubious translation: +lean 3 declaration is + forall {R : Type.{u_2}} [_inst_1 : NonUnitalRing.{u_2} R] (p : RingNorm.{u_2} R (NonUnitalRing.toNonUnitalNonAssocRing.{u_2} R _inst_1)), Eq.{max (succ u_2) 1} (R -> Real) (RingNorm.toFun.{u_2} R (NonUnitalRing.toNonUnitalNonAssocRing.{u_2} R _inst_1) p) (coeFn.{succ u_2, max (succ u_2) 1} (RingNorm.{u_2} R (NonUnitalRing.toNonUnitalNonAssocRing.{u_2} R _inst_1)) (fun (_x : RingNorm.{u_2} R (NonUnitalRing.toNonUnitalNonAssocRing.{u_2} R _inst_1)) => R -> Real) (RingNorm.hasCoeToFun.{u_2} R _inst_1) p) +but is expected to have type + forall {R : Type.{u}} {_inst_1 : Type.{v}}, (Nat -> R -> _inst_1) -> Nat -> (List.{u} R) -> (List.{v} _inst_1) +Case conversion may be inaccurate. Consider using '#align ring_norm.to_fun_eq_coe [anonymous]ₓ'. -/ @[simp] -theorem toFun_eq_coe (p : RingNorm R) : p.toFun = p := +theorem [anonymous] (p : RingNorm R) : p.toFun = p := rfl -#align ring_norm.to_fun_eq_coe RingNorm.toFun_eq_coe - +#align ring_norm.to_fun_eq_coe [anonymous] + +/- warning: ring_norm.ext -> RingNorm.ext is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] {p : RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)} {q : RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)}, (forall (x : R), Eq.{1} Real (coeFn.{succ u1, succ u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (fun (_x : RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) => R -> Real) (RingNorm.hasCoeToFun.{u1} R _inst_1) p x) (coeFn.{succ u1, succ u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (fun (_x : RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) => R -> Real) (RingNorm.hasCoeToFun.{u1} R _inst_1) q x)) -> (Eq.{succ u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) p q) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonUnitalRing.{u1} R] {p : RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)} {q : RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)}, (forall (x : R), Eq.{1} Real (AddGroupSeminorm.toFun.{u1} R (AddCommGroup.toAddGroup.{u1} R (NonUnitalNonAssocRing.toAddCommGroup.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1))) (RingSeminorm.toAddGroupSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) (RingNorm.toRingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) p)) x) (AddGroupSeminorm.toFun.{u1} R (AddCommGroup.toAddGroup.{u1} R (NonUnitalNonAssocRing.toAddCommGroup.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1))) (RingSeminorm.toAddGroupSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) (RingNorm.toRingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) q)) x)) -> (Eq.{succ u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) p q) +Case conversion may be inaccurate. Consider using '#align ring_norm.ext RingNorm.extₓ'. -/ @[ext] theorem ext {p q : RingNorm R} : (∀ x, p x = q x) → p = q := FunLike.ext p q @@ -202,6 +265,12 @@ variable (R) instance [DecidableEq R] : One (RingNorm R) := ⟨{ (1 : RingSeminorm R), (1 : AddGroupNorm R) with }⟩ +/- warning: ring_norm.apply_one -> RingNorm.apply_one is a dubious translation: +lean 3 declaration is + forall (R : Type.{u1}) [_inst_1 : NonUnitalRing.{u1} R] [_inst_2 : DecidableEq.{succ u1} R] (x : R), Eq.{1} Real (coeFn.{succ u1, succ u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (fun (_x : RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) => R -> Real) (RingNorm.hasCoeToFun.{u1} R _inst_1) (OfNat.ofNat.{u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 1 (OfNat.mk.{u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 1 (One.one.{u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (RingNorm.hasOne.{u1} R _inst_1 (fun (a : R) (b : R) => _inst_2 a b))))) x) (ite.{1} Real (Eq.{succ u1} R x (OfNat.ofNat.{u1} R 0 (OfNat.mk.{u1} R 0 (Zero.zero.{u1} R (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))))))) (_inst_2 x (OfNat.ofNat.{u1} R 0 (OfNat.mk.{u1} R 0 (Zero.zero.{u1} R (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))))))) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) +but is expected to have type + forall (R : Type.{u1}) [_inst_1 : NonUnitalRing.{u1} R] [_inst_2 : DecidableEq.{succ u1} R] (x : R), Eq.{1} Real (AddGroupSeminorm.toFun.{u1} R (AddCommGroup.toAddGroup.{u1} R (NonUnitalNonAssocRing.toAddCommGroup.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1))) (RingSeminorm.toAddGroupSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) (RingNorm.toRingSeminorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1) (OfNat.ofNat.{u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) 1 (One.toOfNat1.{u1} (RingNorm.{u1} R (NonUnitalRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (RingNorm.instOneRingNormToNonUnitalNonAssocRing.{u1} R _inst_1 (fun (a : R) (b : R) => _inst_2 a b)))))) x) (ite.{1} Real (Eq.{succ u1} R x (OfNat.ofNat.{u1} R 0 (Zero.toOfNat0.{u1} R (SemigroupWithZero.toZero.{u1} R (NonUnitalSemiring.toSemigroupWithZero.{u1} R (NonUnitalRing.toNonUnitalSemiring.{u1} R _inst_1)))))) (_inst_2 x (OfNat.ofNat.{u1} R 0 (Zero.toOfNat0.{u1} R (SemigroupWithZero.toZero.{u1} R (NonUnitalSemiring.toSemigroupWithZero.{u1} R (NonUnitalRing.toNonUnitalSemiring.{u1} R _inst_1)))))) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) +Case conversion may be inaccurate. Consider using '#align ring_norm.apply_one RingNorm.apply_oneₓ'. -/ @[simp] theorem apply_one [DecidableEq R] (x : R) : (1 : RingNorm R) x = if x = 0 then 0 else 1 := rfl @@ -216,6 +285,7 @@ namespace MulRingSeminorm variable [NonAssocRing R] +#print MulRingSeminorm.mulRingSeminormClass /- instance mulRingSeminormClass : MulRingSeminormClass (MulRingSeminorm R) R ℝ where coe f := f.toFun @@ -226,16 +296,29 @@ instance mulRingSeminormClass : MulRingSeminormClass (MulRingSeminorm R) R ℝ map_mul f := f.map_mul' map_neg_eq_map f := f.neg' #align mul_ring_seminorm.mul_ring_seminorm_class MulRingSeminorm.mulRingSeminormClass +-/ /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ instance : CoeFun (MulRingSeminorm R) fun _ => R → ℝ := FunLike.hasCoeToFun +/- warning: mul_ring_seminorm.to_fun_eq_coe -> MulRingSeminorm.toFun_eq_coe is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonAssocRing.{u1} R] (p : MulRingSeminorm.{u1} R _inst_1), Eq.{succ u1} (R -> Real) (MulRingSeminorm.toFun.{u1} R _inst_1 p) (coeFn.{succ u1, succ u1} (MulRingSeminorm.{u1} R _inst_1) (fun (_x : MulRingSeminorm.{u1} R _inst_1) => R -> Real) (MulRingSeminorm.hasCoeToFun.{u1} R _inst_1) p) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonAssocRing.{u1} R] (p : MulRingSeminorm.{u1} R _inst_1), Eq.{succ u1} (R -> Real) (FunLike.coe.{succ u1, succ u1, 1} (AddGroupSeminorm.{u1} R (AddGroupWithOne.toAddGroup.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (NonAssocRing.toAddCommGroupWithOne.{u1} R _inst_1)))) R (fun (a._@.Mathlib.Analysis.Normed.Group.Seminorm._hyg.956 : R) => Real) (SubadditiveHomClass.toFunLike.{u1, u1, 0} (AddGroupSeminorm.{u1} R (AddGroupWithOne.toAddGroup.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (NonAssocRing.toAddCommGroupWithOne.{u1} R _inst_1)))) R Real (AddZeroClass.toAdd.{u1} R (AddMonoid.toAddZeroClass.{u1} R (SubNegMonoid.toAddMonoid.{u1} R (AddGroup.toSubNegMonoid.{u1} R (AddGroupWithOne.toAddGroup.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (NonAssocRing.toAddCommGroupWithOne.{u1} R _inst_1))))))) (AddZeroClass.toAdd.{0} Real (AddMonoid.toAddZeroClass.{0} Real (AddCommMonoid.toAddMonoid.{0} Real (OrderedAddCommMonoid.toAddCommMonoid.{0} Real Real.orderedAddCommMonoid)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedAddCommMonoid.toPartialOrder.{0} Real Real.orderedAddCommMonoid))) (AddGroupSeminormClass.toSubadditiveHomClass.{u1, u1, 0} (AddGroupSeminorm.{u1} R (AddGroupWithOne.toAddGroup.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (NonAssocRing.toAddCommGroupWithOne.{u1} R _inst_1)))) R Real (AddGroupWithOne.toAddGroup.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (NonAssocRing.toAddCommGroupWithOne.{u1} R _inst_1))) Real.orderedAddCommMonoid (AddGroupSeminorm.addGroupSeminormClass.{u1} R (AddGroupWithOne.toAddGroup.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (NonAssocRing.toAddCommGroupWithOne.{u1} R _inst_1)))))) (MulRingSeminorm.toAddGroupSeminorm.{u1} R _inst_1 p)) (FunLike.coe.{succ u1, succ u1, 1} (MulRingSeminorm.{u1} R _inst_1) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (MulRingSeminormClass.toRingSeminormClass.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real _inst_1 Real.orderedSemiring (MulRingSeminorm.mulRingSeminormClass.{u1} R _inst_1)))) p) +Case conversion may be inaccurate. Consider using '#align mul_ring_seminorm.to_fun_eq_coe MulRingSeminorm.toFun_eq_coeₓ'. -/ @[simp] theorem toFun_eq_coe (p : MulRingSeminorm R) : p.toFun = p := rfl #align mul_ring_seminorm.to_fun_eq_coe MulRingSeminorm.toFun_eq_coe +/- warning: mul_ring_seminorm.ext -> MulRingSeminorm.ext is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonAssocRing.{u1} R] {p : MulRingSeminorm.{u1} R _inst_1} {q : MulRingSeminorm.{u1} R _inst_1}, (forall (x : R), Eq.{1} Real (coeFn.{succ u1, succ u1} (MulRingSeminorm.{u1} R _inst_1) (fun (_x : MulRingSeminorm.{u1} R _inst_1) => R -> Real) (MulRingSeminorm.hasCoeToFun.{u1} R _inst_1) p x) (coeFn.{succ u1, succ u1} (MulRingSeminorm.{u1} R _inst_1) (fun (_x : MulRingSeminorm.{u1} R _inst_1) => R -> Real) (MulRingSeminorm.hasCoeToFun.{u1} R _inst_1) q x)) -> (Eq.{succ u1} (MulRingSeminorm.{u1} R _inst_1) p q) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonAssocRing.{u1} R] {p : MulRingSeminorm.{u1} R _inst_1} {q : MulRingSeminorm.{u1} R _inst_1}, (forall (x : R), Eq.{1} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) x) (FunLike.coe.{succ u1, succ u1, 1} (MulRingSeminorm.{u1} R _inst_1) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (MulRingSeminormClass.toRingSeminormClass.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real _inst_1 Real.orderedSemiring (MulRingSeminorm.mulRingSeminormClass.{u1} R _inst_1)))) p x) (FunLike.coe.{succ u1, succ u1, 1} (MulRingSeminorm.{u1} R _inst_1) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (MulRingSeminormClass.toRingSeminormClass.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real _inst_1 Real.orderedSemiring (MulRingSeminorm.mulRingSeminormClass.{u1} R _inst_1)))) q x)) -> (Eq.{succ u1} (MulRingSeminorm.{u1} R _inst_1) p q) +Case conversion may be inaccurate. Consider using '#align mul_ring_seminorm.ext MulRingSeminorm.extₓ'. -/ @[ext] theorem ext {p q : MulRingSeminorm R} : (∀ x, p x = q x) → p = q := FunLike.ext p q @@ -255,6 +338,12 @@ instance : One (MulRingSeminorm R) := · simp · simp [hx, hy] }⟩ +/- warning: mul_ring_seminorm.apply_one -> MulRingSeminorm.apply_one is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonAssocRing.{u1} R] [_inst_2 : DecidableEq.{succ u1} R] [_inst_3 : NoZeroDivisors.{u1} R (Distrib.toHasMul.{u1} R (NonUnitalNonAssocSemiring.toDistrib.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))) (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1))))] [_inst_4 : Nontrivial.{u1} R] (x : R), Eq.{1} Real (coeFn.{succ u1, succ u1} (MulRingSeminorm.{u1} R _inst_1) (fun (_x : MulRingSeminorm.{u1} R _inst_1) => R -> Real) (MulRingSeminorm.hasCoeToFun.{u1} R _inst_1) (OfNat.ofNat.{u1} (MulRingSeminorm.{u1} R _inst_1) 1 (OfNat.mk.{u1} (MulRingSeminorm.{u1} R _inst_1) 1 (One.one.{u1} (MulRingSeminorm.{u1} R _inst_1) (MulRingSeminorm.hasOne.{u1} R _inst_1 (fun (a : R) (b : R) => _inst_2 a b) _inst_3 _inst_4)))) x) (ite.{1} Real (Eq.{succ u1} R x (OfNat.ofNat.{u1} R 0 (OfNat.mk.{u1} R 0 (Zero.zero.{u1} R (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))))))) (_inst_2 x (OfNat.ofNat.{u1} R 0 (OfNat.mk.{u1} R 0 (Zero.zero.{u1} R (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))))))) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonAssocRing.{u1} R] [_inst_2 : DecidableEq.{succ u1} R] [_inst_3 : NoZeroDivisors.{u1} R (NonUnitalNonAssocRing.toMul.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (MulZeroOneClass.toZero.{u1} R (NonAssocSemiring.toMulZeroOneClass.{u1} R (NonAssocRing.toNonAssocSemiring.{u1} R _inst_1)))] [_inst_4 : Nontrivial.{u1} R] (x : R), Eq.{1} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) x) (FunLike.coe.{succ u1, succ u1, 1} (MulRingSeminorm.{u1} R _inst_1) R (fun (_x : R) => (fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) _x) (SubmultiplicativeHomClass.toFunLike.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real (NonUnitalNonAssocRing.toMul.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (NonUnitalNonAssocSemiring.toMul.{0} Real (NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} Real (Semiring.toNonAssocSemiring.{0} Real (OrderedSemiring.toSemiring.{0} Real Real.orderedSemiring)))) (Preorder.toLE.{0} Real (PartialOrder.toPreorder.{0} Real (OrderedSemiring.toPartialOrder.{0} Real Real.orderedSemiring))) (RingSeminormClass.toSubmultiplicativeHomClass.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1) Real.orderedSemiring (MulRingSeminormClass.toRingSeminormClass.{u1, u1, 0} (MulRingSeminorm.{u1} R _inst_1) R Real _inst_1 Real.orderedSemiring (MulRingSeminorm.mulRingSeminormClass.{u1} R _inst_1)))) (OfNat.ofNat.{u1} (MulRingSeminorm.{u1} R _inst_1) 1 (One.toOfNat1.{u1} (MulRingSeminorm.{u1} R _inst_1) (MulRingSeminorm.instOneMulRingSeminorm.{u1} R _inst_1 (fun (a : R) (b : R) => _inst_2 a b) _inst_3 _inst_4))) x) (ite.{1} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) x) (Eq.{succ u1} R x (OfNat.ofNat.{u1} R 0 (Zero.toOfNat0.{u1} R (MulZeroOneClass.toZero.{u1} R (NonAssocSemiring.toMulZeroOneClass.{u1} R (NonAssocRing.toNonAssocSemiring.{u1} R _inst_1)))))) (_inst_2 x (OfNat.ofNat.{u1} R 0 (Zero.toOfNat0.{u1} R (MulZeroOneClass.toZero.{u1} R (NonAssocSemiring.toMulZeroOneClass.{u1} R (NonAssocRing.toNonAssocSemiring.{u1} R _inst_1)))))) (OfNat.ofNat.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) x) 0 (Zero.toOfNat0.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) x) Real.instZeroReal)) (OfNat.ofNat.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) x) 1 (One.toOfNat1.{0} ((fun (a._@.Mathlib.Analysis.Normed.Ring.Seminorm._hyg.1434 : R) => Real) x) Real.instOneReal))) +Case conversion may be inaccurate. Consider using '#align mul_ring_seminorm.apply_one MulRingSeminorm.apply_oneₓ'. -/ @[simp] theorem apply_one (x : R) : (1 : MulRingSeminorm R) x = if x = 0 then 0 else 1 := rfl @@ -269,6 +358,7 @@ namespace MulRingNorm variable [NonAssocRing R] +#print MulRingNorm.mulRingNormClass /- instance mulRingNormClass : MulRingNormClass (MulRingNorm R) R ℝ where coe f := f.toFun @@ -280,16 +370,30 @@ instance mulRingNormClass : MulRingNormClass (MulRingNorm R) R ℝ map_neg_eq_map f := f.neg' eq_zero_of_map_eq_zero f := f.eq_zero_of_map_eq_zero' #align mul_ring_norm.mul_ring_norm_class MulRingNorm.mulRingNormClass +-/ /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ instance : CoeFun (MulRingNorm R) fun _ => R → ℝ := ⟨fun p => p.toFun⟩ +/- warning: mul_ring_norm.to_fun_eq_coe clashes with [anonymous] -> [anonymous] +warning: mul_ring_norm.to_fun_eq_coe -> [anonymous] is a dubious translation: +lean 3 declaration is + forall {R : Type.{u_2}} [_inst_1 : NonAssocRing.{u_2} R] (p : MulRingNorm.{u_2} R _inst_1), Eq.{max (succ u_2) 1} (R -> Real) (MulRingNorm.toFun.{u_2} R _inst_1 p) (coeFn.{succ u_2, max (succ u_2) 1} (MulRingNorm.{u_2} R _inst_1) (fun (_x : MulRingNorm.{u_2} R _inst_1) => R -> Real) (MulRingNorm.hasCoeToFun.{u_2} R _inst_1) p) +but is expected to have type + forall {R : Type.{u}} {_inst_1 : Type.{v}}, (Nat -> R -> _inst_1) -> Nat -> (List.{u} R) -> (List.{v} _inst_1) +Case conversion may be inaccurate. Consider using '#align mul_ring_norm.to_fun_eq_coe [anonymous]ₓ'. -/ @[simp] -theorem toFun_eq_coe (p : MulRingNorm R) : p.toFun = p := +theorem [anonymous] (p : MulRingNorm R) : p.toFun = p := rfl -#align mul_ring_norm.to_fun_eq_coe MulRingNorm.toFun_eq_coe - +#align mul_ring_norm.to_fun_eq_coe [anonymous] + +/- warning: mul_ring_norm.ext -> MulRingNorm.ext is a dubious translation: +lean 3 declaration is + forall {R : Type.{u1}} [_inst_1 : NonAssocRing.{u1} R] {p : MulRingNorm.{u1} R _inst_1} {q : MulRingNorm.{u1} R _inst_1}, (forall (x : R), Eq.{1} Real (coeFn.{succ u1, succ u1} (MulRingNorm.{u1} R _inst_1) (fun (_x : MulRingNorm.{u1} R _inst_1) => R -> Real) (MulRingNorm.hasCoeToFun.{u1} R _inst_1) p x) (coeFn.{succ u1, succ u1} (MulRingNorm.{u1} R _inst_1) (fun (_x : MulRingNorm.{u1} R _inst_1) => R -> Real) (MulRingNorm.hasCoeToFun.{u1} R _inst_1) q x)) -> (Eq.{succ u1} (MulRingNorm.{u1} R _inst_1) p q) +but is expected to have type + forall {R : Type.{u1}} [_inst_1 : NonAssocRing.{u1} R] {p : MulRingNorm.{u1} R _inst_1} {q : MulRingNorm.{u1} R _inst_1}, (forall (x : R), Eq.{1} Real (AddGroupSeminorm.toFun.{u1} R (AddGroupWithOne.toAddGroup.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (NonAssocRing.toAddCommGroupWithOne.{u1} R _inst_1))) (MulRingSeminorm.toAddGroupSeminorm.{u1} R _inst_1 (MulRingNorm.toMulRingSeminorm.{u1} R _inst_1 p)) x) (AddGroupSeminorm.toFun.{u1} R (AddGroupWithOne.toAddGroup.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (NonAssocRing.toAddCommGroupWithOne.{u1} R _inst_1))) (MulRingSeminorm.toAddGroupSeminorm.{u1} R _inst_1 (MulRingNorm.toMulRingSeminorm.{u1} R _inst_1 q)) x)) -> (Eq.{succ u1} (MulRingNorm.{u1} R _inst_1) p q) +Case conversion may be inaccurate. Consider using '#align mul_ring_norm.ext MulRingNorm.extₓ'. -/ @[ext] theorem ext {p q : MulRingNorm R} : (∀ x, p x = q x) → p = q := FunLike.ext p q @@ -302,6 +406,12 @@ other element. -/ instance : One (MulRingNorm R) := ⟨{ (1 : MulRingSeminorm R), (1 : AddGroupNorm R) with }⟩ +/- warning: mul_ring_norm.apply_one -> MulRingNorm.apply_one is a dubious translation: +lean 3 declaration is + forall (R : Type.{u1}) [_inst_1 : NonAssocRing.{u1} R] [_inst_2 : DecidableEq.{succ u1} R] [_inst_3 : NoZeroDivisors.{u1} R (Distrib.toHasMul.{u1} R (NonUnitalNonAssocSemiring.toDistrib.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))) (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1))))] [_inst_4 : Nontrivial.{u1} R] (x : R), Eq.{1} Real (coeFn.{succ u1, succ u1} (MulRingNorm.{u1} R _inst_1) (fun (_x : MulRingNorm.{u1} R _inst_1) => R -> Real) (MulRingNorm.hasCoeToFun.{u1} R _inst_1) (OfNat.ofNat.{u1} (MulRingNorm.{u1} R _inst_1) 1 (OfNat.mk.{u1} (MulRingNorm.{u1} R _inst_1) 1 (One.one.{u1} (MulRingNorm.{u1} R _inst_1) (MulRingNorm.hasOne.{u1} R _inst_1 (fun (a : R) (b : R) => _inst_2 a b) _inst_3 _inst_4)))) x) (ite.{1} Real (Eq.{succ u1} R x (OfNat.ofNat.{u1} R 0 (OfNat.mk.{u1} R 0 (Zero.zero.{u1} R (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))))))) (_inst_2 x (OfNat.ofNat.{u1} R 0 (OfNat.mk.{u1} R 0 (Zero.zero.{u1} R (MulZeroClass.toHasZero.{u1} R (NonUnitalNonAssocSemiring.toMulZeroClass.{u1} R (NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)))))))) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) +but is expected to have type + forall (R : Type.{u1}) [_inst_1 : NonAssocRing.{u1} R] [_inst_2 : DecidableEq.{succ u1} R] [_inst_3 : NoZeroDivisors.{u1} R (NonUnitalNonAssocRing.toMul.{u1} R (NonAssocRing.toNonUnitalNonAssocRing.{u1} R _inst_1)) (MulZeroOneClass.toZero.{u1} R (NonAssocSemiring.toMulZeroOneClass.{u1} R (NonAssocRing.toNonAssocSemiring.{u1} R _inst_1)))] [_inst_4 : Nontrivial.{u1} R] (x : R), Eq.{1} Real (AddGroupSeminorm.toFun.{u1} R (AddGroupWithOne.toAddGroup.{u1} R (AddCommGroupWithOne.toAddGroupWithOne.{u1} R (NonAssocRing.toAddCommGroupWithOne.{u1} R _inst_1))) (MulRingSeminorm.toAddGroupSeminorm.{u1} R _inst_1 (MulRingNorm.toMulRingSeminorm.{u1} R _inst_1 (OfNat.ofNat.{u1} (MulRingNorm.{u1} R _inst_1) 1 (One.toOfNat1.{u1} (MulRingNorm.{u1} R _inst_1) (MulRingNorm.instOneMulRingNorm.{u1} R _inst_1 (fun (a : R) (b : R) => _inst_2 a b) _inst_3 _inst_4))))) x) (ite.{1} Real (Eq.{succ u1} R x (OfNat.ofNat.{u1} R 0 (Zero.toOfNat0.{u1} R (MulZeroOneClass.toZero.{u1} R (NonAssocSemiring.toMulZeroOneClass.{u1} R (NonAssocRing.toNonAssocSemiring.{u1} R _inst_1)))))) (_inst_2 x (OfNat.ofNat.{u1} R 0 (Zero.toOfNat0.{u1} R (MulZeroOneClass.toZero.{u1} R (NonAssocSemiring.toMulZeroOneClass.{u1} R (NonAssocRing.toNonAssocSemiring.{u1} R _inst_1)))))) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) +Case conversion may be inaccurate. Consider using '#align mul_ring_norm.apply_one MulRingNorm.apply_oneₓ'. -/ @[simp] theorem apply_one (x : R) : (1 : MulRingNorm R) x = if x = 0 then 0 else 1 := rfl @@ -312,6 +422,12 @@ instance : Inhabited (MulRingNorm R) := end MulRingNorm +/- warning: ring_seminorm.to_ring_norm -> RingSeminorm.toRingNorm is a dubious translation: +lean 3 declaration is + forall {K : Type.{u1}} [_inst_1 : Field.{u1} K] (f : RingSeminorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))), (Ne.{succ u1} (RingSeminorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))) f (OfNat.ofNat.{u1} (RingSeminorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))) 0 (OfNat.mk.{u1} (RingSeminorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))) 0 (Zero.zero.{u1} (RingSeminorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))) (RingSeminorm.hasZero.{u1} K (NonUnitalCommRing.toNonUnitalRing.{u1} K (CommRing.toNonUnitalCommRing.{u1} K (EuclideanDomain.toCommRing.{u1} K (Field.toEuclideanDomain.{u1} K _inst_1))))))))) -> (RingNorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))) +but is expected to have type + forall {K : Type.{u1}} [_inst_1 : Field.{u1} K] (f : RingSeminorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))), (Ne.{succ u1} (RingSeminorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))) f (OfNat.ofNat.{u1} (RingSeminorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))) 0 (Zero.toOfNat0.{u1} (RingSeminorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))) (RingSeminorm.instZeroRingSeminormToNonUnitalNonAssocRing.{u1} K (NonUnitalCommRing.toNonUnitalRing.{u1} K (CommRing.toNonUnitalCommRing.{u1} K (EuclideanDomain.toCommRing.{u1} K (Field.toEuclideanDomain.{u1} K _inst_1)))))))) -> (RingNorm.{u1} K (NonAssocRing.toNonUnitalNonAssocRing.{u1} K (Ring.toNonAssocRing.{u1} K (DivisionRing.toRing.{u1} K (Field.toDivisionRing.{u1} K _inst_1))))) +Case conversion may be inaccurate. Consider using '#align ring_seminorm.to_ring_norm RingSeminorm.toRingNormₓ'. -/ /-- A nonzero ring seminorm on a field `K` is a ring norm. -/ def RingSeminorm.toRingNorm {K : Type _} [Field K] (f : RingSeminorm K) (hnt : f ≠ 0) : RingNorm K := @@ -331,9 +447,11 @@ def RingSeminorm.toRingNorm {K : Type _} [Field K] (f : RingSeminorm K) (hnt : f exact hc hc0 } #align ring_seminorm.to_ring_norm RingSeminorm.toRingNorm +#print normRingNorm /- /-- The norm of a normed_ring as a ring_norm. -/ @[simps] def normRingNorm (R : Type _) [NonUnitalNormedRing R] : RingNorm R := { normAddGroupNorm R, normRingSeminorm R with } #align norm_ring_norm normRingNorm +-/ diff --git a/Mathbin/CategoryTheory/Sites/Whiskering.lean b/Mathbin/CategoryTheory/Sites/Whiskering.lean index d37d994668..9357c325a7 100644 --- a/Mathbin/CategoryTheory/Sites/Whiskering.lean +++ b/Mathbin/CategoryTheory/Sites/Whiskering.lean @@ -48,6 +48,7 @@ namespace GrothendieckTopology.Cover variable (P : Cᵒᵖ ⥤ A) {X : C} (S : J.cover X) +#print CategoryTheory.GrothendieckTopology.Cover.multicospanComp /- /-- The multicospan associated to a cover `S : J.cover X` and a presheaf of the form `P ⋙ F` is isomorphic to the composition of the multicospan associated to `S` and `P`, composed with `F`. -/ @@ -62,43 +63,86 @@ def multicospanComp : (S.index (P ⋙ F)).multicospan ≅ (S.index P).multicospa any_goals dsimp; erw [Functor.map_id, Functor.map_id, category.id_comp] any_goals dsimp; erw [category.comp_id, category.id_comp]; rfl) #align category_theory.grothendieck_topology.cover.multicospan_comp CategoryTheory.GrothendieckTopology.Cover.multicospanComp +-/ +/- warning: category_theory.grothendieck_topology.cover.multicospan_comp_app_left -> CategoryTheory.GrothendieckTopology.Cover.multicospanComp_app_left is a dubious translation: +lean 3 declaration is + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u1 u2, max u1 u2, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u1 u2, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (a : CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{succ (max u1 u2)} (CategoryTheory.Iso.{max u1 u2, u4} B _inst_3 (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a))) (CategoryTheory.Iso.app.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.eqToIso.{max u1 u2, u4} B _inst_3 (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (rfl.{succ u4} B (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)))) +but is expected to have type + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u2 u1, max u2 u1, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u2 u1, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (a : CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{max (succ u2) (succ u1)} (CategoryTheory.Iso.{max u2 u1, u4} B _inst_3 (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a))) (CategoryTheory.Iso.app.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.eqToIso.{max u2 u1, u4} B _inst_3 (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (rfl.{succ u4} B (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)))) +Case conversion may be inaccurate. Consider using '#align category_theory.grothendieck_topology.cover.multicospan_comp_app_left CategoryTheory.GrothendieckTopology.Cover.multicospanComp_app_leftₓ'. -/ @[simp] theorem multicospanComp_app_left (a) : (S.multicospanComp F P).app (WalkingMulticospan.left a) = eqToIso rfl := rfl #align category_theory.grothendieck_topology.cover.multicospan_comp_app_left CategoryTheory.GrothendieckTopology.Cover.multicospanComp_app_left +/- warning: category_theory.grothendieck_topology.cover.multicospan_comp_app_right -> CategoryTheory.GrothendieckTopology.Cover.multicospanComp_app_right is a dubious translation: +lean 3 declaration is + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u1 u2, max u1 u2, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u1 u2, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (b : CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{succ (max u1 u2)} (CategoryTheory.Iso.{max u1 u2, u4} B _inst_3 (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b))) (CategoryTheory.Iso.app.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.eqToIso.{max u1 u2, u4} B _inst_3 (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (rfl.{succ u4} B (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)))) +but is expected to have type + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u2 u1, max u2 u1, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u2 u1, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (b : CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{max (succ u2) (succ u1)} (CategoryTheory.Iso.{max u2 u1, u4} B _inst_3 (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b))) (CategoryTheory.Iso.app.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.eqToIso.{max u2 u1, u4} B _inst_3 (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (rfl.{succ u4} B (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)))) +Case conversion may be inaccurate. Consider using '#align category_theory.grothendieck_topology.cover.multicospan_comp_app_right CategoryTheory.GrothendieckTopology.Cover.multicospanComp_app_rightₓ'. -/ @[simp] theorem multicospanComp_app_right (b) : (S.multicospanComp F P).app (WalkingMulticospan.right b) = eqToIso rfl := rfl #align category_theory.grothendieck_topology.cover.multicospan_comp_app_right CategoryTheory.GrothendieckTopology.Cover.multicospanComp_app_right +/- warning: category_theory.grothendieck_topology.cover.multicospan_comp_hom_app_left -> CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_app_left is a dubious translation: +lean 3 declaration is + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u1 u2, max u1 u2, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u1 u2, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (a : CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{succ (max u1 u2)} (Quiver.Hom.{succ (max u1 u2), u4} B (CategoryTheory.CategoryStruct.toQuiver.{max u1 u2, u4} B (CategoryTheory.Category.toCategoryStruct.{max u1 u2, u4} B _inst_3)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a))) (CategoryTheory.NatTrans.app.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Iso.hom.{max (max u2 u1) u1 u2, max (max u2 u1) (max u1 u2) (max u2 u1) u4} (CategoryTheory.Functor.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.eqToHom.{max u1 u2, u4} B _inst_3 (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (rfl.{succ u4} B (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)))) +but is expected to have type + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u2 u1, max u2 u1, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u2 u1, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (a : CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{max (succ u2) (succ u1)} (Quiver.Hom.{succ (max u2 u1), u4} B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a))) (CategoryTheory.NatTrans.app.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Iso.hom.{max u2 u1, max (max u2 u4) u1} (CategoryTheory.Functor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.eqToHom.{max u2 u1, u4} B _inst_3 (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (rfl.{succ u4} B (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)))) +Case conversion may be inaccurate. Consider using '#align category_theory.grothendieck_topology.cover.multicospan_comp_hom_app_left CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_app_leftₓ'. -/ @[simp] theorem multicospanComp_hom_app_left (a) : (S.multicospanComp F P).Hom.app (WalkingMulticospan.left a) = eqToHom rfl := rfl #align category_theory.grothendieck_topology.cover.multicospan_comp_hom_app_left CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_app_left +/- warning: category_theory.grothendieck_topology.cover.multicospan_comp_hom_app_right -> CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_app_right is a dubious translation: +lean 3 declaration is + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u1 u2, max u1 u2, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u1 u2, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (b : CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{succ (max u1 u2)} (Quiver.Hom.{succ (max u1 u2), u4} B (CategoryTheory.CategoryStruct.toQuiver.{max u1 u2, u4} B (CategoryTheory.Category.toCategoryStruct.{max u1 u2, u4} B _inst_3)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b))) (CategoryTheory.NatTrans.app.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Iso.hom.{max (max u2 u1) u1 u2, max (max u2 u1) (max u1 u2) (max u2 u1) u4} (CategoryTheory.Functor.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.eqToHom.{max u1 u2, u4} B _inst_3 (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (rfl.{succ u4} B (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)))) +but is expected to have type + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u2 u1, max u2 u1, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u2 u1, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (b : CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{max (succ u2) (succ u1)} (Quiver.Hom.{succ (max u2 u1), u4} B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b))) (CategoryTheory.NatTrans.app.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Iso.hom.{max u2 u1, max (max u2 u4) u1} (CategoryTheory.Functor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.eqToHom.{max u2 u1, u4} B _inst_3 (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (rfl.{succ u4} B (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)))) +Case conversion may be inaccurate. Consider using '#align category_theory.grothendieck_topology.cover.multicospan_comp_hom_app_right CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_app_rightₓ'. -/ @[simp] theorem multicospanComp_hom_app_right (b) : (S.multicospanComp F P).Hom.app (WalkingMulticospan.right b) = eqToHom rfl := rfl #align category_theory.grothendieck_topology.cover.multicospan_comp_hom_app_right CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_app_right +/- warning: category_theory.grothendieck_topology.cover.multicospan_comp_hom_inv_left -> CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_inv_left is a dubious translation: +lean 3 declaration is + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u1 u2, max u1 u2, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u1 u2, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (a : CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{succ (max u1 u2)} (Quiver.Hom.{succ (max u1 u2), u4} B (CategoryTheory.CategoryStruct.toQuiver.{max u1 u2, u4} B (CategoryTheory.Category.toCategoryStruct.{max u1 u2, u4} B _inst_3)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a))) (CategoryTheory.NatTrans.app.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Iso.inv.{max (max u2 u1) u1 u2, max (max u2 u1) (max u1 u2) (max u2 u1) u4} (CategoryTheory.Functor.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.eqToHom.{max u1 u2, u4} B _inst_3 (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (rfl.{succ u4} B (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)))) +but is expected to have type + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u2 u1, max u2 u1, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u2 u1, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (a : CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{max (succ u2) (succ u1)} (Quiver.Hom.{succ (max u2 u1), u4} B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a))) (CategoryTheory.NatTrans.app.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Iso.inv.{max u2 u1, max (max u2 u4) u1} (CategoryTheory.Functor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (CategoryTheory.eqToHom.{max u2 u1, u4} B _inst_3 (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)) (rfl.{succ u4} B (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.left.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) a)))) +Case conversion may be inaccurate. Consider using '#align category_theory.grothendieck_topology.cover.multicospan_comp_hom_inv_left CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_inv_leftₓ'. -/ @[simp] theorem multicospanComp_hom_inv_left (P : Cᵒᵖ ⥤ A) {X : C} (S : J.cover X) (a) : (S.multicospanComp F P).inv.app (WalkingMulticospan.left a) = eqToHom rfl := rfl #align category_theory.grothendieck_topology.cover.multicospan_comp_hom_inv_left CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_inv_left +/- warning: category_theory.grothendieck_topology.cover.multicospan_comp_hom_inv_right -> CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_inv_right is a dubious translation: +lean 3 declaration is + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u1 u2, max u1 u2, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u1 u2, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (b : CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{succ (max u1 u2)} (Quiver.Hom.{succ (max u1 u2), u4} B (CategoryTheory.CategoryStruct.toQuiver.{max u1 u2, u4} B (CategoryTheory.Category.toCategoryStruct.{max u1 u2, u4} B _inst_3)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b))) (CategoryTheory.NatTrans.app.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Iso.inv.{max (max u2 u1) u1 u2, max (max u2 u1) (max u1 u2) (max u2 u1) u4} (CategoryTheory.Functor.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.eqToHom.{max u1 u2, u4} B _inst_3 (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (rfl.{succ u4} B (CategoryTheory.Functor.obj.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)))) +but is expected to have type + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u2 u1, max u2 u1, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u2 u1, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X) (b : CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))), Eq.{max (succ u2) (succ u1)} (Quiver.Hom.{succ (max u2 u1), u4} B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b))) (CategoryTheory.NatTrans.app.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Iso.inv.{max u2 u1, max (max u2 u4) u1} (CategoryTheory.Functor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (CategoryTheory.eqToHom.{max u2 u1, u4} B _inst_3 (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)) (rfl.{succ u4} B (Prefunctor.obj.{succ (max u2 u1), succ (max u2 u1), max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max u2 u1} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) B (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, u4} B (CategoryTheory.Category.toCategoryStruct.{max u2 u1, u4} B _inst_3)) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.WalkingMulticospan.right.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) b)))) +Case conversion may be inaccurate. Consider using '#align category_theory.grothendieck_topology.cover.multicospan_comp_hom_inv_right CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_inv_rightₓ'. -/ @[simp] theorem multicospanComp_hom_inv_right (P : Cᵒᵖ ⥤ A) {X : C} (S : J.cover X) (b) : (S.multicospanComp F P).inv.app (WalkingMulticospan.right b) = eqToHom rfl := rfl #align category_theory.grothendieck_topology.cover.multicospan_comp_hom_inv_right CategoryTheory.GrothendieckTopology.Cover.multicospanComp_hom_inv_right +/- warning: category_theory.grothendieck_topology.cover.map_multifork -> CategoryTheory.GrothendieckTopology.Cover.mapMultifork is a dubious translation: +lean 3 declaration is + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u1 u2, max u1 u2, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u1 u2, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X), CategoryTheory.Iso.{max u1 u2, max (max u2 u1) u4 u1 u2} (CategoryTheory.Limits.Cone.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.Cone.category.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Functor.mapCone.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F (CategoryTheory.GrothendieckTopology.Cover.multifork.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Functor.obj.{max u1 u2, max u1 u2, max (max u2 u1) u4 u1 u2, max (max u2 u1) u4 u1 u2} (CategoryTheory.Limits.Cone.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.Cone.category.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.Cone.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.Cone.category.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.Cones.postcompose.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Iso.hom.{max (max u2 u1) u1 u2, max (max u2 u1) (max u1 u2) (max u2 u1) u4} (CategoryTheory.Functor.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u1 u2, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u1 u2, max u1 u2, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.CategoryTheory.smallCategory.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u1 u2, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u1 u2, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S))) (CategoryTheory.GrothendieckTopology.Cover.multifork.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u1 u2, max u1 u2, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) +but is expected to have type + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] {A : Type.{u3}} [_inst_2 : CategoryTheory.Category.{max u1 u2, u3} A] {B : Type.{u4}} [_inst_3 : CategoryTheory.Category.{max u1 u2, u4} B] {J : CategoryTheory.GrothendieckTopology.{u1, u2} C _inst_1} (F : CategoryTheory.Functor.{max u2 u1, max u2 u1, u3, u4} A _inst_2 B _inst_3) (P : CategoryTheory.Functor.{u1, max u2 u1, u2, u3} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2) {X : C} (S : CategoryTheory.GrothendieckTopology.Cover.{u1, u2} C _inst_1 J X), CategoryTheory.Iso.{max u2 u1, max (max (max u2 u4) u1) u2 u1} (CategoryTheory.Limits.Cone.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.Cone.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Functor.mapCone.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u3} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P))) A _inst_2 B _inst_3 F (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (CategoryTheory.GrothendieckTopology.Cover.multifork.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) (Prefunctor.obj.{max (succ u2) (succ u1), max (succ u2) (succ u1), max (max u2 u4) u1, max (max u2 u4) u1} (CategoryTheory.Limits.Cone.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max (max u2 u4) u1} (CategoryTheory.Limits.Cone.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max (max u2 u4) u1} (CategoryTheory.Limits.Cone.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.Cone.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))))) (CategoryTheory.Limits.Cone.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.CategoryStruct.toQuiver.{max u2 u1, max (max u2 u4) u1} (CategoryTheory.Limits.Cone.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Category.toCategoryStruct.{max u2 u1, max (max u2 u4) u1} (CategoryTheory.Limits.Cone.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.Cone.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)))) (CategoryTheory.Functor.toPrefunctor.{max u2 u1, max u2 u1, max (max u2 u4) u1, max (max u2 u4) u1} (CategoryTheory.Limits.Cone.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.Cone.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.Cone.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.Cone.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F)) (CategoryTheory.Limits.Cones.postcompose.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.Iso.hom.{max u2 u1, max (max u2 u4) u1} (CategoryTheory.Functor.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Functor.category.{max u2 u1, max u2 u1, max u2 u1, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) B _inst_3) (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u4, max u2 u1} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Functor.comp.{max u2 u1, max u2 u1, max u2 u1, max u2 u1, u3, u4} (CategoryTheory.Limits.WalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) (CategoryTheory.Limits.WalkingMulticospan.instSmallCategoryWalkingMulticospan.{max u2 u1} (CategoryTheory.Limits.MulticospanIndex.L.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.R.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.fstTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) (CategoryTheory.Limits.MulticospanIndex.sndTo.{max u2 u1, max u2 u1, u4} B _inst_3 (CategoryTheory.GrothendieckTopology.Cover.index.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F)))) A _inst_2 B _inst_3 (CategoryTheory.Limits.MulticospanIndex.multicospan.{max u2 u1, u3, max u2 u1} A _inst_2 (CategoryTheory.GrothendieckTopology.Cover.index.{u3, u1, u2} C _inst_1 X J A _inst_2 S P)) F) (CategoryTheory.GrothendieckTopology.Cover.multicospanComp.{u1, u2, u3, u4} C _inst_1 A _inst_2 B _inst_3 J F P X S)))) (CategoryTheory.GrothendieckTopology.Cover.multifork.{u4, u1, u2} C _inst_1 X J B _inst_3 S (CategoryTheory.Functor.comp.{u1, max u2 u1, max u2 u1, u2, u3, u4} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) A _inst_2 B _inst_3 P F))) +Case conversion may be inaccurate. Consider using '#align category_theory.grothendieck_topology.cover.map_multifork CategoryTheory.GrothendieckTopology.Cover.mapMultiforkₓ'. -/ /-- Mapping the multifork associated to a cover `S : J.cover X` and a presheaf `P` with respect to a functor `F` is isomorphic (upto a natural isomorphism of the underlying functors) to the multifork associated to `S` and `P ⋙ F`. -/ @@ -120,6 +164,7 @@ end GrothendieckTopology.Cover variable [∀ (X : C) (S : J.cover X) (P : Cᵒᵖ ⥤ A), PreservesLimit (S.index P).multicospan F] +#print CategoryTheory.Presheaf.IsSheaf.comp /- theorem Presheaf.IsSheaf.comp {P : Cᵒᵖ ⥤ A} (hP : Presheaf.IsSheaf J P) : Presheaf.IsSheaf J (P ⋙ F) := by @@ -130,9 +175,11 @@ theorem Presheaf.IsSheaf.comp {P : Cᵒᵖ ⥤ A} (hP : Presheaf.IsSheaf J P) : replace h := limits.is_limit.of_iso_limit h (S.map_multifork F P) exact ⟨limits.is_limit.postcompose_hom_equiv (S.multicospan_comp F P) _ h⟩ #align category_theory.presheaf.is_sheaf.comp CategoryTheory.Presheaf.IsSheaf.comp +-/ variable (J) +#print CategoryTheory.sheafCompose /- /-- Composing a sheaf with a functor preserving the appropriate limits yields a functor between sheaf categories. -/ @[simps] @@ -143,6 +190,7 @@ def sheafCompose : Sheaf J A ⥤ Sheaf J B map_id' G := Sheaf.Hom.ext _ _ <| whiskerRight_id _ map_comp' G H W f g := Sheaf.Hom.ext _ _ <| whiskerRight_comp _ _ _ #align category_theory.Sheaf_compose CategoryTheory.sheafCompose +-/ end CategoryTheory diff --git a/Mathbin/Data/Int/Cast/Lemmas.lean b/Mathbin/Data/Int/Cast/Lemmas.lean index 1c41314ef3..f08adad18a 100644 --- a/Mathbin/Data/Int/Cast/Lemmas.lean +++ b/Mathbin/Data/Int/Cast/Lemmas.lean @@ -625,7 +625,7 @@ end Pi lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u2}} {γ : Type.{u3}} [_inst_1 : IntCast.{u3} γ] (n : Int), Eq.{max (max (succ u1) (succ u2)) (succ u3)} ((Sum.{u1, u2} α β) -> γ) (Sum.elim.{u1, u2, succ u3} α β γ ((fun (a : Type) (b : Sort.{max (succ u1) (succ u3)}) [self : HasLiftT.{1, max (succ u1) (succ u3)} a b] => self.0) Int (α -> γ) (HasLiftT.mk.{1, max (succ u1) (succ u3)} Int (α -> γ) (CoeTCₓ.coe.{1, max (succ u1) (succ u3)} Int (α -> γ) (Int.castCoe.{max u1 u3} (α -> γ) (Pi.hasIntCast.{u1, u3} α (fun (ᾰ : α) => γ) (fun (i : α) => _inst_1))))) n) ((fun (a : Type) (b : Sort.{max (succ u2) (succ u3)}) [self : HasLiftT.{1, max (succ u2) (succ u3)} a b] => self.0) Int (β -> γ) (HasLiftT.mk.{1, max (succ u2) (succ u3)} Int (β -> γ) (CoeTCₓ.coe.{1, max (succ u2) (succ u3)} Int (β -> γ) (Int.castCoe.{max u2 u3} (β -> γ) (Pi.hasIntCast.{u2, u3} β (fun (ᾰ : β) => γ) (fun (i : β) => _inst_1))))) n)) ((fun (a : Type) (b : Sort.{max (max (succ u1) (succ u2)) (succ u3)}) [self : HasLiftT.{1, max (max (succ u1) (succ u2)) (succ u3)} a b] => self.0) Int ((Sum.{u1, u2} α β) -> γ) (HasLiftT.mk.{1, max (max (succ u1) (succ u2)) (succ u3)} Int ((Sum.{u1, u2} α β) -> γ) (CoeTCₓ.coe.{1, max (max (succ u1) (succ u2)) (succ u3)} Int ((Sum.{u1, u2} α β) -> γ) (Int.castCoe.{max (max u1 u2) u3} ((Sum.{u1, u2} α β) -> γ) (Pi.hasIntCast.{max u1 u2, u3} (Sum.{u1, u2} α β) (fun (ᾰ : Sum.{u1, u2} α β) => γ) (fun (i : Sum.{u1, u2} α β) => _inst_1))))) n) but is expected to have type - forall {α : Type.{u3}} {β : Type.{u2}} {γ : Type.{u1}} [_inst_1 : IntCast.{u1} γ] (n : Int), Eq.{max (max (succ u3) (succ u2)) (succ u1)} ((Sum.{u3, u2} α β) -> γ) (Sum.elim.{u3, u2, succ u1} α β γ (Int.cast.{max u3 u1} (α -> γ) (Pi.intCast.{u3, u1} α (fun (a._@.Mathlib.Data.Int.Cast.Lemmas._hyg.3011 : α) => γ) (fun (i : α) => _inst_1)) n) (Int.cast.{max u2 u1} (β -> γ) (Pi.intCast.{u2, u1} β (fun (a._@.Mathlib.Data.Int.Cast.Lemmas._hyg.3018 : β) => γ) (fun (i : β) => _inst_1)) n)) (Int.cast.{max (max u3 u2) u1} ((Sum.{u3, u2} α β) -> γ) (Pi.intCast.{max u3 u2, u1} (Sum.{u3, u2} α β) (fun (a._@.Mathlib.Data.Sum.Basic._hyg.1871 : Sum.{u3, u2} α β) => γ) (fun (i : Sum.{u3, u2} α β) => _inst_1)) n) + forall {α : Type.{u3}} {β : Type.{u2}} {γ : Type.{u1}} [_inst_1 : IntCast.{u1} γ] (n : Int), Eq.{max (max (succ u3) (succ u2)) (succ u1)} ((Sum.{u3, u2} α β) -> γ) (Sum.elim.{u3, u2, succ u1} α β γ (Int.cast.{max u3 u1} (α -> γ) (Pi.intCast.{u3, u1} α (fun (a._@.Mathlib.Data.Int.Cast.Lemmas._hyg.3012 : α) => γ) (fun (i : α) => _inst_1)) n) (Int.cast.{max u2 u1} (β -> γ) (Pi.intCast.{u2, u1} β (fun (a._@.Mathlib.Data.Int.Cast.Lemmas._hyg.3019 : β) => γ) (fun (i : β) => _inst_1)) n)) (Int.cast.{max (max u3 u2) u1} ((Sum.{u3, u2} α β) -> γ) (Pi.intCast.{max u3 u2, u1} (Sum.{u3, u2} α β) (fun (a._@.Mathlib.Data.Sum.Basic._hyg.1871 : Sum.{u3, u2} α β) => γ) (fun (i : Sum.{u3, u2} α β) => _inst_1)) n) Case conversion may be inaccurate. Consider using '#align sum.elim_int_cast_int_cast Sum.elim_intCast_intCastₓ'. -/ theorem Sum.elim_intCast_intCast {α β γ : Type _} [IntCast γ] (n : ℤ) : Sum.elim (n : α → γ) (n : β → γ) = n := diff --git a/Mathbin/GroupTheory/FreeAbelianGroup.lean b/Mathbin/GroupTheory/FreeAbelianGroup.lean index 88d663bd69..dff87cb9df 100644 --- a/Mathbin/GroupTheory/FreeAbelianGroup.lean +++ b/Mathbin/GroupTheory/FreeAbelianGroup.lean @@ -209,7 +209,7 @@ protected theorem induction_on {C : FreeAbelianGroup α → Prop} (z : FreeAbeli lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : AddCommGroup.{u2} β] (a : FreeAbelianGroup.{u1} α) (f : α -> β) (g : α -> β), Eq.{succ u2} β (coeFn.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) (fun (_x : AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) => (FreeAbelianGroup.{u1} α) -> β) (AddMonoidHom.hasCoeToFun.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) (coeFn.{max 1 (max (max (succ u1) (succ u2)) (succ u2) (succ u1)) (max (succ u2) (succ u1)) (succ u1) (succ u2), max (max (succ u1) (succ u2)) (succ u2) (succ u1)} (Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (fun (_x : Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) => (α -> β) -> (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (Equiv.hasCoeToFun.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (FreeAbelianGroup.lift.{u1, u2} α β _inst_1) (HAdd.hAdd.{max u1 u2, max u1 u2, max u1 u2} (α -> β) (α -> β) (α -> β) (instHAdd.{max u1 u2} (α -> β) (Pi.instAdd.{u1, u2} α (fun (ᾰ : α) => β) (fun (i : α) => AddZeroClass.toHasAdd.{u2} β (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))))) f g)) a) (HAdd.hAdd.{u2, u2, u2} β β β (instHAdd.{u2} β (AddZeroClass.toHasAdd.{u2} β (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (coeFn.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) (fun (_x : AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) => (FreeAbelianGroup.{u1} α) -> β) (AddMonoidHom.hasCoeToFun.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) (coeFn.{max 1 (max (max (succ u1) (succ u2)) (succ u2) (succ u1)) (max (succ u2) (succ u1)) (succ u1) (succ u2), max (max (succ u1) (succ u2)) (succ u2) (succ u1)} (Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (fun (_x : Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) => (α -> β) -> (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (Equiv.hasCoeToFun.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (FreeAbelianGroup.lift.{u1, u2} α β _inst_1) f) a) (coeFn.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) (fun (_x : AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) => (FreeAbelianGroup.{u1} α) -> β) (AddMonoidHom.hasCoeToFun.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) (coeFn.{max 1 (max (max (succ u1) (succ u2)) (succ u2) (succ u1)) (max (succ u2) (succ u1)) (succ u1) (succ u2), max (max (succ u1) (succ u2)) (succ u2) (succ u1)} (Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (fun (_x : Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) => (α -> β) -> (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (Equiv.hasCoeToFun.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (FreeAbelianGroup.lift.{u1, u2} α β _inst_1) g) a)) but is expected to have type - forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : AddCommGroup.{u1} β] (a : FreeAbelianGroup.{u2} α) (f : α -> β) (g : α -> β), Eq.{succ u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (HAdd.hAdd.{max u1 u2, max u1 u2, max u1 u2} (α -> β) (α -> β) (α -> β) (instHAdd.{max u1 u2} (α -> β) (Pi.instAdd.{u2, u1} α (fun (a._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.784 : α) => β) (fun (i : α) => AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) f g)) (FreeAbelianGroup.{u2} α) (fun (_x : FreeAbelianGroup.{u2} α) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) _x) (AddHomClass.toFunLike.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (HAdd.hAdd.{max u1 u2, max u1 u2, max u1 u2} (α -> β) (α -> β) (α -> β) (instHAdd.{max u1 u2} (α -> β) (Pi.instAdd.{u2, u1} α (fun (a._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.784 : α) => β) (fun (i : α) => AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) f g)) (FreeAbelianGroup.{u2} α) β (AddZeroClass.toAdd.{u2} (FreeAbelianGroup.{u2} α) (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α)))))) (AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (AddMonoidHomClass.toAddHomClass.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (HAdd.hAdd.{max u1 u2, max u1 u2, max u1 u2} (α -> β) (α -> β) (α -> β) (instHAdd.{max u1 u2} (α -> β) (Pi.instAdd.{u2, u1} α (fun (a._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.784 : α) => β) (fun (i : α) => AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) f g)) (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))) (AddMonoidHom.addMonoidHomClass.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) (HAdd.hAdd.{max u1 u2, max u1 u2, max u1 u2} (α -> β) (α -> β) (α -> β) (instHAdd.{max u1 u2} (α -> β) (Pi.instAdd.{u2, u1} α (fun (ᾰ : α) => β) (fun (i : α) => AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) f g)) a) (HAdd.hAdd.{u1, u1, u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (instHAdd.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (AddZeroClass.toAdd.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (AddMonoid.toAddZeroClass.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (SubNegMonoid.toAddMonoid.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (AddGroup.toSubNegMonoid.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (AddCommGroup.toAddGroup.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) _inst_1)))))) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) f) (FreeAbelianGroup.{u2} α) (fun (_x : FreeAbelianGroup.{u2} α) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) _x) (AddHomClass.toFunLike.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) f) (FreeAbelianGroup.{u2} α) β (AddZeroClass.toAdd.{u2} (FreeAbelianGroup.{u2} α) (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α)))))) (AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (AddMonoidHomClass.toAddHomClass.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) f) (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))) (AddMonoidHom.addMonoidHomClass.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) f) a) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) g) (FreeAbelianGroup.{u2} α) (fun (_x : FreeAbelianGroup.{u2} α) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) _x) (AddHomClass.toFunLike.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) g) (FreeAbelianGroup.{u2} α) β (AddZeroClass.toAdd.{u2} (FreeAbelianGroup.{u2} α) (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α)))))) (AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (AddMonoidHomClass.toAddHomClass.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) g) (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))) (AddMonoidHom.addMonoidHomClass.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) g) a)) + forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : AddCommGroup.{u1} β] (a : FreeAbelianGroup.{u2} α) (f : α -> β) (g : α -> β), Eq.{succ u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (HAdd.hAdd.{max u1 u2, max u1 u2, max u1 u2} (α -> β) (α -> β) (α -> β) (instHAdd.{max u1 u2} (α -> β) (Pi.instAdd.{u2, u1} α (fun (a._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.785 : α) => β) (fun (i : α) => AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) f g)) (FreeAbelianGroup.{u2} α) (fun (_x : FreeAbelianGroup.{u2} α) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) _x) (AddHomClass.toFunLike.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (HAdd.hAdd.{max u1 u2, max u1 u2, max u1 u2} (α -> β) (α -> β) (α -> β) (instHAdd.{max u1 u2} (α -> β) (Pi.instAdd.{u2, u1} α (fun (a._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.785 : α) => β) (fun (i : α) => AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) f g)) (FreeAbelianGroup.{u2} α) β (AddZeroClass.toAdd.{u2} (FreeAbelianGroup.{u2} α) (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α)))))) (AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (AddMonoidHomClass.toAddHomClass.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (HAdd.hAdd.{max u1 u2, max u1 u2, max u1 u2} (α -> β) (α -> β) (α -> β) (instHAdd.{max u1 u2} (α -> β) (Pi.instAdd.{u2, u1} α (fun (a._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.785 : α) => β) (fun (i : α) => AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) f g)) (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))) (AddMonoidHom.addMonoidHomClass.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) (HAdd.hAdd.{max u1 u2, max u1 u2, max u1 u2} (α -> β) (α -> β) (α -> β) (instHAdd.{max u1 u2} (α -> β) (Pi.instAdd.{u2, u1} α (fun (ᾰ : α) => β) (fun (i : α) => AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) f g)) a) (HAdd.hAdd.{u1, u1, u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (instHAdd.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (AddZeroClass.toAdd.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (AddMonoid.toAddZeroClass.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (SubNegMonoid.toAddMonoid.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (AddGroup.toSubNegMonoid.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) (AddCommGroup.toAddGroup.{u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) a) _inst_1)))))) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) f) (FreeAbelianGroup.{u2} α) (fun (_x : FreeAbelianGroup.{u2} α) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) _x) (AddHomClass.toFunLike.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) f) (FreeAbelianGroup.{u2} α) β (AddZeroClass.toAdd.{u2} (FreeAbelianGroup.{u2} α) (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α)))))) (AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (AddMonoidHomClass.toAddHomClass.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) f) (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))) (AddMonoidHom.addMonoidHomClass.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) f) a) (FunLike.coe.{max (succ u1) (succ u2), succ u2, succ u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) g) (FreeAbelianGroup.{u2} α) (fun (_x : FreeAbelianGroup.{u2} α) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u2} α) => β) _x) (AddHomClass.toFunLike.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) g) (FreeAbelianGroup.{u2} α) β (AddZeroClass.toAdd.{u2} (FreeAbelianGroup.{u2} α) (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α)))))) (AddZeroClass.toAdd.{u1} β (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (AddMonoidHomClass.toAddHomClass.{max u1 u2, u2, u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) g) (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))) (AddMonoidHom.addMonoidHomClass.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))))) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) g) a)) Case conversion may be inaccurate. Consider using '#align free_abelian_group.lift.add' FreeAbelianGroup.lift.add'ₓ'. -/ theorem lift.add' {α β} [AddCommGroup β] (a : FreeAbelianGroup α) (f g : α → β) : lift (f + g) a = lift f a + lift g a := @@ -238,7 +238,7 @@ def liftAddGroupHom {α} (β) [AddCommGroup β] (a : FreeAbelianGroup α) : (α lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : AddCommGroup.{u2} β] (f : α -> β), Eq.{max (succ u2) (succ u1)} (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) (coeFn.{max 1 (max (max (succ u1) (succ u2)) (succ u2) (succ u1)) (max (succ u2) (succ u1)) (succ u1) (succ u2), max (max (succ u1) (succ u2)) (succ u2) (succ u1)} (Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (fun (_x : Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) => (α -> β) -> (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (Equiv.hasCoeToFun.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (FreeAbelianGroup.lift.{u1, u2} α β _inst_1) (Neg.neg.{max u1 u2} (α -> β) (Pi.instNeg.{u1, u2} α (fun (ᾰ : α) => β) (fun (i : α) => SubNegMonoid.toHasNeg.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))) f)) (Neg.neg.{max u2 u1} (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1))))) (AddMonoidHom.hasNeg.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) _inst_1) (coeFn.{max 1 (max (max (succ u1) (succ u2)) (succ u2) (succ u1)) (max (succ u2) (succ u1)) (succ u1) (succ u2), max (max (succ u1) (succ u2)) (succ u2) (succ u1)} (Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (fun (_x : Equiv.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) => (α -> β) -> (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (Equiv.hasCoeToFun.{max (succ u1) (succ u2), max (succ u2) (succ u1)} (α -> β) (AddMonoidHom.{u1, u2} (FreeAbelianGroup.{u1} α) β (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) (AddMonoid.toAddZeroClass.{u2} β (SubNegMonoid.toAddMonoid.{u2} β (AddGroup.toSubNegMonoid.{u2} β (AddCommGroup.toAddGroup.{u2} β _inst_1)))))) (FreeAbelianGroup.lift.{u1, u2} α β _inst_1) f)) but is expected to have type - forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : AddCommGroup.{u1} β] (f : α -> β), Eq.{max (succ u2) (succ u1)} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (Neg.neg.{max u2 u1} (α -> β) (Pi.instNeg.{u2, u1} α (fun (a._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.909 : α) => β) (fun (i : α) => NegZeroClass.toNeg.{u1} β (SubNegZeroMonoid.toNegZeroClass.{u1} β (SubtractionMonoid.toSubNegZeroMonoid.{u1} β (SubtractionCommMonoid.toSubtractionMonoid.{u1} β (AddCommGroup.toDivisionAddCommMonoid.{u1} β _inst_1)))))) f)) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) (Neg.neg.{max u2 u1} (α -> β) (Pi.instNeg.{u2, u1} α (fun (ᾰ : α) => β) (fun (i : α) => NegZeroClass.toNeg.{u1} β (SubNegZeroMonoid.toNegZeroClass.{u1} β (SubtractionMonoid.toSubNegZeroMonoid.{u1} β (SubtractionCommMonoid.toSubtractionMonoid.{u1} β (AddCommGroup.toDivisionAddCommMonoid.{u1} β _inst_1)))))) f)) (Neg.neg.{max u2 u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) f) (AddMonoidHom.instNegAddMonoidHomToAddZeroClassToAddMonoidToSubNegAddMonoidToAddGroup.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) _inst_1) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) f)) + forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : AddCommGroup.{u1} β] (f : α -> β), Eq.{max (succ u2) (succ u1)} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) (Neg.neg.{max u2 u1} (α -> β) (Pi.instNeg.{u2, u1} α (fun (a._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.910 : α) => β) (fun (i : α) => NegZeroClass.toNeg.{u1} β (SubNegZeroMonoid.toNegZeroClass.{u1} β (SubtractionMonoid.toSubNegZeroMonoid.{u1} β (SubtractionCommMonoid.toSubtractionMonoid.{u1} β (AddCommGroup.toDivisionAddCommMonoid.{u1} β _inst_1)))))) f)) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) (Neg.neg.{max u2 u1} (α -> β) (Pi.instNeg.{u2, u1} α (fun (ᾰ : α) => β) (fun (i : α) => NegZeroClass.toNeg.{u1} β (SubNegZeroMonoid.toNegZeroClass.{u1} β (SubtractionMonoid.toSubNegZeroMonoid.{u1} β (SubtractionCommMonoid.toSubtractionMonoid.{u1} β (AddCommGroup.toDivisionAddCommMonoid.{u1} β _inst_1)))))) f)) (Neg.neg.{max u2 u1} ((fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) f) (AddMonoidHom.instNegAddMonoidHomToAddZeroClassToAddMonoidToSubNegAddMonoidToAddGroup.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) _inst_1) (FunLike.coe.{max (succ u1) (succ u2), max (succ u1) (succ u2), max (succ u1) (succ u2)} (Equiv.{max (succ u2) (succ u1), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (α -> β) (fun (_x : α -> β) => (fun (x._@.Mathlib.Logic.Equiv.Defs._hyg.808 : α -> β) => AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1))))) _x) (Equiv.instFunLikeEquiv.{max (succ u1) (succ u2), max (succ u1) (succ u2)} (α -> β) (AddMonoidHom.{u2, u1} (FreeAbelianGroup.{u2} α) β (AddMonoid.toAddZeroClass.{u2} (FreeAbelianGroup.{u2} α) (SubNegMonoid.toAddMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddGroup.toSubNegMonoid.{u2} (FreeAbelianGroup.{u2} α) (AddCommGroup.toAddGroup.{u2} (FreeAbelianGroup.{u2} α) (FreeAbelianGroup.addCommGroup.{u2} α))))) (AddMonoid.toAddZeroClass.{u1} β (SubNegMonoid.toAddMonoid.{u1} β (AddGroup.toSubNegMonoid.{u1} β (AddCommGroup.toAddGroup.{u1} β _inst_1)))))) (FreeAbelianGroup.lift.{u2, u1} α β _inst_1) f)) Case conversion may be inaccurate. Consider using '#align free_abelian_group.lift_neg' FreeAbelianGroup.lift_neg'ₓ'. -/ theorem lift_neg' {β} [AddCommGroup β] (f : α → β) : lift (-f) = -lift f := AddMonoidHom.ext fun _ => (liftAddGroupHom _ _ : (α → β) →+ β).map_neg _ @@ -394,7 +394,7 @@ theorem sub_bind (f : α → FreeAbelianGroup β) (x y : FreeAbelianGroup α) : lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u1}} (f : α -> β) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} (fun {α : Type.{u1}} => FreeAbelianGroup.{u1} α) (Applicative.toHasSeq.{u1, u1} (fun {α : Type.{u1}} => FreeAbelianGroup.{u1} α) (Monad.toApplicative.{u1, u1} (fun {α : Type.{u1}} => FreeAbelianGroup.{u1} α) FreeAbelianGroup.monad.{u1})) α β (Pure.pure.{u1, u1} (fun {α : Type.{u1}} => FreeAbelianGroup.{u1} α) (Applicative.toHasPure.{u1, u1} (fun {α : Type.{u1}} => FreeAbelianGroup.{u1} α) (Monad.toApplicative.{u1, u1} (fun {α : Type.{u1}} => FreeAbelianGroup.{u1} α) FreeAbelianGroup.monad.{u1})) (α -> β) f) x) (Functor.map.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toFunctor.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f x) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u1}} (f : α -> β) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (Pure.pure.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toPure.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) (α -> β) f) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1667 : Unit) => x)) (Functor.map.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toFunctor.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f x) + forall {α : Type.{u1}} {β : Type.{u1}} (f : α -> β) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (Pure.pure.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toPure.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) (α -> β) f) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1668 : Unit) => x)) (Functor.map.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toFunctor.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f x) Case conversion may be inaccurate. Consider using '#align free_abelian_group.pure_seq FreeAbelianGroup.pure_seqₓ'. -/ @[simp] theorem pure_seq (f : α → β) (x : FreeAbelianGroup α) : pure f <*> x = f <$> x := @@ -405,7 +405,7 @@ theorem pure_seq (f : α → β) (x : FreeAbelianGroup α) : pure f <*> x = f <$ lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u1}} (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} (α -> β)) 0 (OfNat.mk.{u1} (FreeAbelianGroup.{u1} (α -> β)) 0 (Zero.zero.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddZeroClass.toHasZero.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β)))))))))) x) (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} β) 0 (OfNat.mk.{u1} (FreeAbelianGroup.{u1} β) 0 (Zero.zero.{u1} (FreeAbelianGroup.{u1} β) (AddZeroClass.toHasZero.{u1} (FreeAbelianGroup.{u1} β) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))))))) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u1}} (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} (α -> β)) 0 (Zero.toOfNat0.{u1} (FreeAbelianGroup.{u1} (α -> β)) (NegZeroClass.toZero.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β))))))))) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1701 : Unit) => x)) (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} β) 0 (Zero.toOfNat0.{u1} (FreeAbelianGroup.{u1} β) (NegZeroClass.toZero.{u1} (FreeAbelianGroup.{u1} β) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} β) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))))))) + forall {α : Type.{u1}} {β : Type.{u1}} (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} (α -> β)) 0 (Zero.toOfNat0.{u1} (FreeAbelianGroup.{u1} (α -> β)) (NegZeroClass.toZero.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β))))))))) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1702 : Unit) => x)) (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} β) 0 (Zero.toOfNat0.{u1} (FreeAbelianGroup.{u1} β) (NegZeroClass.toZero.{u1} (FreeAbelianGroup.{u1} β) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} β) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))))))) Case conversion may be inaccurate. Consider using '#align free_abelian_group.zero_seq FreeAbelianGroup.zero_seqₓ'. -/ @[simp] theorem zero_seq (x : FreeAbelianGroup α) : (0 : FreeAbelianGroup (α → β)) <*> x = 0 := @@ -416,7 +416,7 @@ theorem zero_seq (x : FreeAbelianGroup α) : (0 : FreeAbelianGroup (α → β)) lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (g : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (instHAdd.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddZeroClass.toHasAdd.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β)))))))) f g) x) (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHAdd.{u1} (FreeAbelianGroup.{u1} β) (AddZeroClass.toHasAdd.{u1} (FreeAbelianGroup.{u1} β) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f x) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β g x)) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (g : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (instHAdd.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β)))))))) f g) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1739 : Unit) => x)) (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHAdd.{u1} (FreeAbelianGroup.{u1} β) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} β) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1750 : Unit) => x)) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β g (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1760 : Unit) => x))) + forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (g : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (instHAdd.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β)))))))) f g) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1740 : Unit) => x)) (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHAdd.{u1} (FreeAbelianGroup.{u1} β) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} β) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1751 : Unit) => x)) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β g (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1761 : Unit) => x))) Case conversion may be inaccurate. Consider using '#align free_abelian_group.add_seq FreeAbelianGroup.add_seqₓ'. -/ @[simp] theorem add_seq (f g : FreeAbelianGroup (α → β)) (x : FreeAbelianGroup α) : @@ -428,7 +428,7 @@ theorem add_seq (f g : FreeAbelianGroup (α → β)) (x : FreeAbelianGroup α) : lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β (Neg.neg.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegMonoid.toHasNeg.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β))))) f) x) (Neg.neg.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toHasNeg.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f x)) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (Neg.neg.{u1} (FreeAbelianGroup.{u1} (α -> β)) (NegZeroClass.toNeg.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β))))))) f) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1792 : Unit) => x)) (Neg.neg.{u1} (FreeAbelianGroup.{u1} β) (NegZeroClass.toNeg.{u1} (FreeAbelianGroup.{u1} β) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} β) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1803 : Unit) => x))) + forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (Neg.neg.{u1} (FreeAbelianGroup.{u1} (α -> β)) (NegZeroClass.toNeg.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β))))))) f) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1793 : Unit) => x)) (Neg.neg.{u1} (FreeAbelianGroup.{u1} β) (NegZeroClass.toNeg.{u1} (FreeAbelianGroup.{u1} β) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} β) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1804 : Unit) => x))) Case conversion may be inaccurate. Consider using '#align free_abelian_group.neg_seq FreeAbelianGroup.neg_seqₓ'. -/ @[simp] theorem neg_seq (f : FreeAbelianGroup (α → β)) (x : FreeAbelianGroup α) : -f <*> x = -(f <*> x) := @@ -439,7 +439,7 @@ theorem neg_seq (f : FreeAbelianGroup (α → β)) (x : FreeAbelianGroup α) : - lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (g : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (instHSub.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegMonoid.toHasSub.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β)))))) f g) x) (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHSub.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toHasSub.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f x) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β g x)) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (g : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (instHSub.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegMonoid.toSub.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β)))))) f g) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1841 : Unit) => x)) (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHSub.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toSub.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1852 : Unit) => x)) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β g (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1862 : Unit) => x))) + forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (g : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.{u1} (α -> β)) (instHSub.{u1} (FreeAbelianGroup.{u1} (α -> β)) (SubNegMonoid.toSub.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} (α -> β)) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} (α -> β)) (FreeAbelianGroup.addCommGroup.{u1} (α -> β)))))) f g) (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1842 : Unit) => x)) (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHSub.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toSub.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1853 : Unit) => x)) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β g (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.1863 : Unit) => x))) Case conversion may be inaccurate. Consider using '#align free_abelian_group.sub_seq FreeAbelianGroup.sub_seqₓ'. -/ @[simp] theorem sub_seq (f g : FreeAbelianGroup (α → β)) (x : FreeAbelianGroup α) : @@ -463,7 +463,7 @@ def seqAddGroupHom (f : FreeAbelianGroup (α → β)) : FreeAbelianGroup α →+ lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} α) 0 (OfNat.mk.{u1} (FreeAbelianGroup.{u1} α) 0 (Zero.zero.{u1} (FreeAbelianGroup.{u1} α) (AddZeroClass.toHasZero.{u1} (FreeAbelianGroup.{u1} α) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α)))))))))) (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} β) 0 (OfNat.mk.{u1} (FreeAbelianGroup.{u1} β) 0 (Zero.zero.{u1} (FreeAbelianGroup.{u1} β) (AddZeroClass.toHasZero.{u1} (FreeAbelianGroup.{u1} β) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))))))) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2011 : Unit) => OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} α) 0 (Zero.toOfNat0.{u1} (FreeAbelianGroup.{u1} α) (NegZeroClass.toZero.{u1} (FreeAbelianGroup.{u1} α) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} α) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))))))) (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} β) 0 (Zero.toOfNat0.{u1} (FreeAbelianGroup.{u1} β) (NegZeroClass.toZero.{u1} (FreeAbelianGroup.{u1} β) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} β) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))))))) + forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2012 : Unit) => OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} α) 0 (Zero.toOfNat0.{u1} (FreeAbelianGroup.{u1} α) (NegZeroClass.toZero.{u1} (FreeAbelianGroup.{u1} α) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} α) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))))))) (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} β) 0 (Zero.toOfNat0.{u1} (FreeAbelianGroup.{u1} β) (NegZeroClass.toZero.{u1} (FreeAbelianGroup.{u1} β) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} β) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))))))) Case conversion may be inaccurate. Consider using '#align free_abelian_group.seq_zero FreeAbelianGroup.seq_zeroₓ'. -/ @[simp] theorem seq_zero (f : FreeAbelianGroup (α → β)) : f <*> 0 = 0 := @@ -474,7 +474,7 @@ theorem seq_zero (f : FreeAbelianGroup (α → β)) : f <*> 0 = 0 := lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α) (y : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (instHAdd.{u1} (FreeAbelianGroup.{u1} α) (AddZeroClass.toHasAdd.{u1} (FreeAbelianGroup.{u1} α) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))))) x y)) (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHAdd.{u1} (FreeAbelianGroup.{u1} β) (AddZeroClass.toHasAdd.{u1} (FreeAbelianGroup.{u1} β) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f x) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f y)) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α) (y : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2044 : Unit) => HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (instHAdd.{u1} (FreeAbelianGroup.{u1} α) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} α) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))))) x y)) (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHAdd.{u1} (FreeAbelianGroup.{u1} β) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} β) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2059 : Unit) => x)) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2069 : Unit) => y))) + forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α) (y : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2045 : Unit) => HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (instHAdd.{u1} (FreeAbelianGroup.{u1} α) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} α) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))))) x y)) (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHAdd.{u1} (FreeAbelianGroup.{u1} β) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} β) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2060 : Unit) => x)) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2070 : Unit) => y))) Case conversion may be inaccurate. Consider using '#align free_abelian_group.seq_add FreeAbelianGroup.seq_addₓ'. -/ @[simp] theorem seq_add (f : FreeAbelianGroup (α → β)) (x y : FreeAbelianGroup α) : @@ -486,7 +486,7 @@ theorem seq_add (f : FreeAbelianGroup (α → β)) (x y : FreeAbelianGroup α) : lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f (Neg.neg.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toHasNeg.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α)))) x)) (Neg.neg.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toHasNeg.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f x)) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2101 : Unit) => Neg.neg.{u1} (FreeAbelianGroup.{u1} α) (NegZeroClass.toNeg.{u1} (FreeAbelianGroup.{u1} α) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} α) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α)))))) x)) (Neg.neg.{u1} (FreeAbelianGroup.{u1} β) (NegZeroClass.toNeg.{u1} (FreeAbelianGroup.{u1} β) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} β) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2115 : Unit) => x))) + forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2102 : Unit) => Neg.neg.{u1} (FreeAbelianGroup.{u1} α) (NegZeroClass.toNeg.{u1} (FreeAbelianGroup.{u1} α) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} α) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} α) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α)))))) x)) (Neg.neg.{u1} (FreeAbelianGroup.{u1} β) (NegZeroClass.toNeg.{u1} (FreeAbelianGroup.{u1} β) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} β) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} β) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β)))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2116 : Unit) => x))) Case conversion may be inaccurate. Consider using '#align free_abelian_group.seq_neg FreeAbelianGroup.seq_negₓ'. -/ @[simp] theorem seq_neg (f : FreeAbelianGroup (α → β)) (x : FreeAbelianGroup α) : f <*> -x = -(f <*> x) := @@ -497,7 +497,7 @@ theorem seq_neg (f : FreeAbelianGroup (α → β)) (x : FreeAbelianGroup α) : f lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α) (y : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (instHSub.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toHasSub.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) x y)) (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHSub.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toHasSub.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f x) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toHasSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.monad.{u1})) α β f y)) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α) (y : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2148 : Unit) => HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (instHSub.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toSub.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) x y)) (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHSub.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toSub.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2163 : Unit) => x)) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2173 : Unit) => y))) + forall {α : Type.{u1}} {β : Type.{u1}} (f : FreeAbelianGroup.{u1} (α -> β)) (x : FreeAbelianGroup.{u1} α) (y : FreeAbelianGroup.{u1} α), Eq.{succ u1} (FreeAbelianGroup.{u1} β) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2149 : Unit) => HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.{u1} α) (instHSub.{u1} (FreeAbelianGroup.{u1} α) (SubNegMonoid.toSub.{u1} (FreeAbelianGroup.{u1} α) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} α) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α))))) x y)) (HSub.hSub.{u1, u1, u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.{u1} β) (instHSub.{u1} (FreeAbelianGroup.{u1} β) (SubNegMonoid.toSub.{u1} (FreeAbelianGroup.{u1} β) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} β) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} β) (FreeAbelianGroup.addCommGroup.{u1} β))))) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2164 : Unit) => x)) (Seq.seq.{u1, u1} FreeAbelianGroup.{u1} (Applicative.toSeq.{u1, u1} FreeAbelianGroup.{u1} (Monad.toApplicative.{u1, u1} FreeAbelianGroup.{u1} FreeAbelianGroup.instMonadFreeAbelianGroup.{u1})) α β f (fun (x._@.Mathlib.GroupTheory.FreeAbelianGroup._hyg.2174 : Unit) => y))) Case conversion may be inaccurate. Consider using '#align free_abelian_group.seq_sub FreeAbelianGroup.seq_subₓ'. -/ @[simp] theorem seq_sub (f : FreeAbelianGroup (α → β)) (x y : FreeAbelianGroup α) : diff --git a/Mathbin/GroupTheory/FreeAbelianGroupFinsupp.lean b/Mathbin/GroupTheory/FreeAbelianGroupFinsupp.lean index 1a0c7bdd0e..f0998866ad 100644 --- a/Mathbin/GroupTheory/FreeAbelianGroupFinsupp.lean +++ b/Mathbin/GroupTheory/FreeAbelianGroupFinsupp.lean @@ -36,11 +36,23 @@ open BigOperators variable {X : Type _} +/- warning: free_abelian_group.to_finsupp -> FreeAbelianGroup.toFinsupp is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}}, AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) +but is expected to have type + forall {X : Type.{u1}}, AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.to_finsupp FreeAbelianGroup.toFinsuppₓ'. -/ /-- The group homomorphism `free_abelian_group X →+ (X →₀ ℤ)`. -/ def FreeAbelianGroup.toFinsupp : FreeAbelianGroup X →+ X →₀ ℤ := FreeAbelianGroup.lift fun x => Finsupp.single x (1 : ℤ) #align free_abelian_group.to_finsupp FreeAbelianGroup.toFinsupp +/- warning: finsupp.to_free_abelian_group -> Finsupp.toFreeAbelianGroup is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}}, AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int Int.hasZero) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) +but is expected to have type + forall {X : Type.{u1}}, AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) +Case conversion may be inaccurate. Consider using '#align finsupp.to_free_abelian_group Finsupp.toFreeAbelianGroupₓ'. -/ /-- The group homomorphism `(X →₀ ℤ) →+ free_abelian_group X`. -/ def Finsupp.toFreeAbelianGroup : (X →₀ ℤ) →+ FreeAbelianGroup X := Finsupp.liftAddHom fun x => (smulAddHom ℤ (FreeAbelianGroup X)).flip (FreeAbelianGroup.of x) @@ -48,6 +60,12 @@ def Finsupp.toFreeAbelianGroup : (X →₀ ℤ) →+ FreeAbelianGroup X := open Finsupp FreeAbelianGroup +/- warning: finsupp.to_free_abelian_group_comp_single_add_hom -> Finsupp.toFreeAbelianGroup_comp_singleAddHom is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}} (x : X), Eq.{succ u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddMonoidHom.comp.{0, u1, u1} Int (Finsupp.{u1, 0} X Int Int.hasZero) (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.toFreeAbelianGroup.{u1} X) (Finsupp.singleAddHom.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid) x)) (coeFn.{succ u1, succ u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddMonoid.toAddZeroClass.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddCommMonoid.toAddMonoid.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addCommMonoid.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (fun (_x : AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddMonoid.toAddZeroClass.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddCommMonoid.toAddMonoid.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addCommMonoid.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) => (FreeAbelianGroup.{u1} X) -> (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddMonoidHom.hasCoeToFun.{u1, u1} (FreeAbelianGroup.{u1} X) (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddMonoid.toAddZeroClass.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddCommMonoid.toAddMonoid.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addCommMonoid.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddMonoidHom.flip.{0, u1, u1} Int (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)) (smulAddHom.{0, u1} Int (FreeAbelianGroup.{u1} X) Int.semiring (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)) (AddCommGroup.intModule.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (FreeAbelianGroup.of.{u1} X x)) +but is expected to have type + forall {X : Type.{u1}} (x : X), Eq.{succ u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddMonoidHom.comp.{0, u1, u1} Int (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.toFreeAbelianGroup.{u1} X) (Finsupp.singleAddHom.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt) x)) (FunLike.coe.{succ u1, succ u1, succ u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddMonoid.toAddZeroClass.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddCommMonoid.toAddMonoid.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addCommMonoid.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (FreeAbelianGroup.{u1} X) (fun (_x : FreeAbelianGroup.{u1} X) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) _x) (AddHomClass.toFunLike.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddMonoid.toAddZeroClass.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddCommMonoid.toAddMonoid.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addCommMonoid.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (FreeAbelianGroup.{u1} X) (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddZeroClass.toAdd.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddCommMonoid.toAddMonoid.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addCommMonoid.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddMonoidHomClass.toAddHomClass.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddMonoid.toAddZeroClass.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddCommMonoid.toAddMonoid.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addCommMonoid.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (FreeAbelianGroup.{u1} X) (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddMonoid.toAddZeroClass.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddCommMonoid.toAddMonoid.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addCommMonoid.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addMonoidHomClass.{u1, u1} (FreeAbelianGroup.{u1} X) (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddMonoid.toAddZeroClass.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddCommMonoid.toAddMonoid.{u1} (AddMonoidHom.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addCommMonoid.{0, u1} Int (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))))) (AddMonoidHom.flip.{0, u1, u1} Int (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{0} Int (AddMonoidWithOne.toAddMonoid.{0} Int (AddCommMonoidWithOne.toAddMonoidWithOne.{0} Int (NonAssocSemiring.toAddCommMonoidWithOne.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (AddCommMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)) (smulAddHom.{0, u1} Int (FreeAbelianGroup.{u1} X) Int.instSemiringInt (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)) (AddCommGroup.intModule.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) (FreeAbelianGroup.of.{u1} X x)) +Case conversion may be inaccurate. Consider using '#align finsupp.to_free_abelian_group_comp_single_add_hom Finsupp.toFreeAbelianGroup_comp_singleAddHomₓ'. -/ @[simp] theorem Finsupp.toFreeAbelianGroup_comp_singleAddHom (x : X) : Finsupp.toFreeAbelianGroup.comp (Finsupp.singleAddHom x) = @@ -58,6 +76,12 @@ theorem Finsupp.toFreeAbelianGroup_comp_singleAddHom (x : X) : to_free_abelian_group, Finsupp.liftAddHom_apply_single] #align finsupp.to_free_abelian_group_comp_single_add_hom Finsupp.toFreeAbelianGroup_comp_singleAddHom +/- warning: free_abelian_group.to_finsupp_comp_to_free_abelian_group -> FreeAbelianGroup.toFinsupp_comp_toFreeAbelianGroup is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}}, Eq.{succ u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int Int.hasZero) (Finsupp.{u1, 0} X Int Int.hasZero) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) (AddMonoidHom.comp.{u1, u1, u1} (Finsupp.{u1, 0} X Int Int.hasZero) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (FreeAbelianGroup.toFinsupp.{u1} X) (Finsupp.toFreeAbelianGroup.{u1} X)) (AddMonoidHom.id.{u1} (Finsupp.{u1, 0} X Int Int.hasZero) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) +but is expected to have type + forall {X : Type.{u1}}, Eq.{succ u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (AddMonoidHom.comp.{u1, u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (FreeAbelianGroup.toFinsupp.{u1} X) (Finsupp.toFreeAbelianGroup.{u1} X)) (AddMonoidHom.id.{u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.to_finsupp_comp_to_free_abelian_group FreeAbelianGroup.toFinsupp_comp_toFreeAbelianGroupₓ'. -/ @[simp] theorem FreeAbelianGroup.toFinsupp_comp_toFreeAbelianGroup : toFinsupp.comp toFreeAbelianGroup = AddMonoidHom.id (X →₀ ℤ) := @@ -68,6 +92,7 @@ theorem FreeAbelianGroup.toFinsupp_comp_toFreeAbelianGroup : one_smul, lift.of, AddMonoidHom.flip_apply, smulAddHom_apply, AddMonoidHom.id_apply] #align free_abelian_group.to_finsupp_comp_to_free_abelian_group FreeAbelianGroup.toFinsupp_comp_toFreeAbelianGroup +#print Finsupp.toFreeAbelianGroup_comp_toFinsupp /- @[simp] theorem Finsupp.toFreeAbelianGroup_comp_toFinsupp : toFreeAbelianGroup.comp toFinsupp = AddMonoidHom.id (FreeAbelianGroup X) := @@ -77,7 +102,14 @@ theorem Finsupp.toFreeAbelianGroup_comp_toFinsupp : lift_add_hom_apply_single, AddMonoidHom.flip_apply, smulAddHom_apply, one_smul, AddMonoidHom.id_apply] #align finsupp.to_free_abelian_group_comp_to_finsupp Finsupp.toFreeAbelianGroup_comp_toFinsupp +-/ +/- warning: finsupp.to_free_abelian_group_to_finsupp -> Finsupp.toFreeAbelianGroup_toFinsupp is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}} (x : FreeAbelianGroup.{u1} X), Eq.{succ u1} (FreeAbelianGroup.{u1} X) (coeFn.{succ u1, succ u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int Int.hasZero) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (fun (_x : AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int Int.hasZero) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) => (Finsupp.{u1, 0} X Int Int.hasZero) -> (FreeAbelianGroup.{u1} X)) (AddMonoidHom.hasCoeToFun.{u1, u1} (Finsupp.{u1, 0} X Int Int.hasZero) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.toFreeAbelianGroup.{u1} X) (coeFn.{succ u1, succ u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) (fun (_x : AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) => (FreeAbelianGroup.{u1} X) -> (Finsupp.{u1, 0} X Int Int.hasZero)) (AddMonoidHom.hasCoeToFun.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) (FreeAbelianGroup.toFinsupp.{u1} X) x)) x +but is expected to have type + forall {X : Type.{u1}} (x : FreeAbelianGroup.{u1} X), Eq.{succ u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) => FreeAbelianGroup.{u1} X) (FunLike.coe.{succ u1, succ u1, succ u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (fun (a : FreeAbelianGroup.{u1} X) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) a) (AddHomClass.toFunLike.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddZeroClass.toAdd.{u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (AddMonoidHomClass.toAddHomClass.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoidHom.addMonoidHomClass.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))))) (FreeAbelianGroup.toFinsupp.{u1} X) x)) (FunLike.coe.{succ u1, succ u1, succ u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (fun (_x : Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) => FreeAbelianGroup.{u1} X) _x) (AddHomClass.toFunLike.{u1, u1, u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (AddZeroClass.toAdd.{u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddMonoidHomClass.toAddHomClass.{u1, u1, u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addMonoidHomClass.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))))) (Finsupp.toFreeAbelianGroup.{u1} X) (FunLike.coe.{succ u1, succ u1, succ u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (fun (_x : FreeAbelianGroup.{u1} X) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) _x) (AddHomClass.toFunLike.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddZeroClass.toAdd.{u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (AddMonoidHomClass.toAddHomClass.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoidHom.addMonoidHomClass.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))))) (FreeAbelianGroup.toFinsupp.{u1} X) x)) x +Case conversion may be inaccurate. Consider using '#align finsupp.to_free_abelian_group_to_finsupp Finsupp.toFreeAbelianGroup_toFinsuppₓ'. -/ @[simp] theorem Finsupp.toFreeAbelianGroup_toFinsupp {X} (x : FreeAbelianGroup X) : x.toFinsupp.toFreeAbelianGroup = x := by @@ -90,11 +122,23 @@ open Finsupp variable {X} +/- warning: free_abelian_group.to_finsupp_of -> FreeAbelianGroup.toFinsupp_of is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}} (x : X), Eq.{succ u1} (Finsupp.{u1, 0} X Int Int.hasZero) (coeFn.{succ u1, succ u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) (fun (_x : AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) => (FreeAbelianGroup.{u1} X) -> (Finsupp.{u1, 0} X Int Int.hasZero)) (AddMonoidHom.hasCoeToFun.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) (FreeAbelianGroup.toFinsupp.{u1} X) (FreeAbelianGroup.of.{u1} X x)) (Finsupp.single.{u1, 0} X Int Int.hasZero x (OfNat.ofNat.{0} Int 1 (OfNat.mk.{0} Int 1 (One.one.{0} Int Int.hasOne)))) +but is expected to have type + forall {X : Type.{u1}} (x : X), Eq.{succ u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.of.{u1} X x)) (FunLike.coe.{succ u1, succ u1, succ u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (fun (_x : FreeAbelianGroup.{u1} X) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) _x) (AddHomClass.toFunLike.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddZeroClass.toAdd.{u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (AddMonoidHomClass.toAddHomClass.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoidHom.addMonoidHomClass.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))))) (FreeAbelianGroup.toFinsupp.{u1} X) (FreeAbelianGroup.of.{u1} X x)) (Finsupp.single.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing))))) x (OfNat.ofNat.{0} Int 1 (instOfNatInt 1))) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.to_finsupp_of FreeAbelianGroup.toFinsupp_ofₓ'. -/ @[simp] theorem toFinsupp_of (x : X) : toFinsupp (of x) = Finsupp.single x 1 := by simp only [to_finsupp, lift.of] #align free_abelian_group.to_finsupp_of FreeAbelianGroup.toFinsupp_of +/- warning: free_abelian_group.to_finsupp_to_free_abelian_group -> FreeAbelianGroup.toFinsupp_toFreeAbelianGroup is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}} (f : Finsupp.{u1, 0} X Int Int.hasZero), Eq.{succ u1} (Finsupp.{u1, 0} X Int Int.hasZero) (coeFn.{succ u1, succ u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) (fun (_x : AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) => (FreeAbelianGroup.{u1} X) -> (Finsupp.{u1, 0} X Int Int.hasZero)) (AddMonoidHom.hasCoeToFun.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid))) (FreeAbelianGroup.toFinsupp.{u1} X) (coeFn.{succ u1, succ u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int Int.hasZero) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (fun (_x : AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int Int.hasZero) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) => (Finsupp.{u1, 0} X Int Int.hasZero) -> (FreeAbelianGroup.{u1} X)) (AddMonoidHom.hasCoeToFun.{u1, u1} (Finsupp.{u1, 0} X Int Int.hasZero) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.toFreeAbelianGroup.{u1} X) f)) f +but is expected to have type + forall {X : Type.{u1}} (f : Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))), Eq.{succ u1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FunLike.coe.{succ u1, succ u1, succ u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (fun (a : Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) => FreeAbelianGroup.{u1} X) a) (AddHomClass.toFunLike.{u1, u1, u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (AddZeroClass.toAdd.{u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddMonoidHomClass.toAddHomClass.{u1, u1, u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addMonoidHomClass.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))))) (Finsupp.toFreeAbelianGroup.{u1} X) f)) (FunLike.coe.{succ u1, succ u1, succ u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (fun (_x : FreeAbelianGroup.{u1} X) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) _x) (AddHomClass.toFunLike.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddZeroClass.toAdd.{u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (AddMonoidHomClass.toAddHomClass.{u1, u1, u1} (AddMonoidHom.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoidHom.addMonoidHomClass.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))))) (FreeAbelianGroup.toFinsupp.{u1} X) (FunLike.coe.{succ u1, succ u1, succ u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (fun (_x : Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) => FreeAbelianGroup.{u1} X) _x) (AddHomClass.toFunLike.{u1, u1, u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (AddZeroClass.toAdd.{u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt))) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddMonoidHomClass.toAddHomClass.{u1, u1, u1} (AddMonoidHom.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoidHom.addMonoidHomClass.{u1, u1} (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (FreeAbelianGroup.{u1} X) (Finsupp.addZeroClass.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))))) (Finsupp.toFreeAbelianGroup.{u1} X) f)) f +Case conversion may be inaccurate. Consider using '#align free_abelian_group.to_finsupp_to_free_abelian_group FreeAbelianGroup.toFinsupp_toFreeAbelianGroupₓ'. -/ @[simp] theorem toFinsupp_toFreeAbelianGroup (f : X →₀ ℤ) : f.toFreeAbelianGroup.toFinsupp = f := by rw [← AddMonoidHom.comp_apply, to_finsupp_comp_to_free_abelian_group, AddMonoidHom.id_apply] @@ -102,6 +146,12 @@ theorem toFinsupp_toFreeAbelianGroup (f : X →₀ ℤ) : f.toFreeAbelianGroup.t variable (X) +/- warning: free_abelian_group.equiv_finsupp -> FreeAbelianGroup.equivFinsupp is a dubious translation: +lean 3 declaration is + forall (X : Type.{u1}), AddEquiv.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int Int.hasZero) (AddZeroClass.toHasAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.add.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) +but is expected to have type + forall (X : Type.{u1}), AddEquiv.{u1, u1} (FreeAbelianGroup.{u1} X) (Finsupp.{u1, 0} X Int (CommMonoidWithZero.toZero.{0} Int (CancelCommMonoidWithZero.toCommMonoidWithZero.{0} Int (IsDomain.toCancelCommMonoidWithZero.{0} Int Int.instCommSemiringInt (LinearOrderedRing.isDomain.{0} Int (LinearOrderedCommRing.toLinearOrderedRing.{0} Int Int.linearOrderedCommRing)))))) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (Finsupp.add.{u1, 0} X Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.equiv_finsupp FreeAbelianGroup.equivFinsuppₓ'. -/ /-- The additive equivalence between `free_abelian_group X` and `(X →₀ ℤ)`. -/ @[simps] def equivFinsupp : FreeAbelianGroup X ≃+ (X →₀ ℤ) @@ -113,11 +163,23 @@ def equivFinsupp : FreeAbelianGroup X ≃+ (X →₀ ℤ) map_add' := toFinsupp.map_add #align free_abelian_group.equiv_finsupp FreeAbelianGroup.equivFinsupp +/- warning: free_abelian_group.basis -> FreeAbelianGroup.basis is a dubious translation: +lean 3 declaration is + forall (α : Type.{u1}), Basis.{u1, 0, u1} α Int (FreeAbelianGroup.{u1} α) Int.semiring (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α)) (AddCommGroup.intModule.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α)) +but is expected to have type + forall (α : Type.{u1}), Basis.{u1, 0, u1} α Int (FreeAbelianGroup.{u1} α) Int.instSemiringInt (AddCommGroup.toAddCommMonoid.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α)) (AddCommGroup.intModule.{u1} (FreeAbelianGroup.{u1} α) (FreeAbelianGroup.addCommGroup.{u1} α)) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.basis FreeAbelianGroup.basisₓ'. -/ /-- `A` is a basis of the ℤ-module `free_abelian_group A`. -/ noncomputable def basis (α : Type _) : Basis α ℤ (FreeAbelianGroup α) := ⟨(FreeAbelianGroup.equivFinsupp α).toIntLinearEquiv⟩ #align free_abelian_group.basis FreeAbelianGroup.basis +/- warning: free_abelian_group.equiv.of_free_abelian_group_linear_equiv -> FreeAbelianGroup.Equiv.ofFreeAbelianGroupLinearEquiv is a dubious translation: +lean 3 declaration is + forall {α : Type.{u_1}} {β : Type.{u_2}}, (LinearEquiv.{0, 0, u_1, u_2} Int Int Int.semiring Int.semiring (RingHom.id.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring)) (RingHom.id.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.semiring)) (RingHomInvPair.ids.{0} Int Int.semiring) (RingHomInvPair.ids.{0} Int Int.semiring) (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.{u_2} β) (AddCommGroup.toAddCommMonoid.{u_1} (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.addCommGroup.{u_1} α)) (AddCommGroup.toAddCommMonoid.{u_2} (FreeAbelianGroup.{u_2} β) (FreeAbelianGroup.addCommGroup.{u_2} β)) (AddCommGroup.intModule.{u_1} (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.addCommGroup.{u_1} α)) (AddCommGroup.intModule.{u_2} (FreeAbelianGroup.{u_2} β) (FreeAbelianGroup.addCommGroup.{u_2} β))) -> (Equiv.{succ u_1, succ u_2} α β) +but is expected to have type + forall {α : Type.{u_1}} {β : Type.{u_2}}, (LinearEquiv.{0, 0, u_1, u_2} Int Int Int.instSemiringInt Int.instSemiringInt (RingHom.id.{0} Int (NonAssocRing.toNonAssocSemiring.{0} Int (Ring.toNonAssocRing.{0} Int Int.instRingInt))) (RingHom.id.{0} Int (Semiring.toNonAssocSemiring.{0} Int Int.instSemiringInt)) (RingHomInvPair.ids.{0} Int Int.instSemiringInt) (RingHomInvPair.ids.{0} Int Int.instSemiringInt) (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.{u_2} β) (AddCommGroup.toAddCommMonoid.{u_1} (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.addCommGroup.{u_1} α)) (AddCommGroup.toAddCommMonoid.{u_2} (FreeAbelianGroup.{u_2} β) (FreeAbelianGroup.addCommGroup.{u_2} β)) (AddCommGroup.intModule.{u_1} (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.addCommGroup.{u_1} α)) (AddCommGroup.intModule.{u_2} (FreeAbelianGroup.{u_2} β) (FreeAbelianGroup.addCommGroup.{u_2} β))) -> (Equiv.{succ u_1, succ u_2} α β) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.equiv.of_free_abelian_group_linear_equiv FreeAbelianGroup.Equiv.ofFreeAbelianGroupLinearEquivₓ'. -/ /-- Isomorphic free ablian groups (as modules) have equivalent bases. -/ def Equiv.ofFreeAbelianGroupLinearEquiv {α β : Type _} (e : FreeAbelianGroup α ≃ₗ[ℤ] FreeAbelianGroup β) : α ≃ β := @@ -125,12 +187,24 @@ def Equiv.ofFreeAbelianGroupLinearEquiv {α β : Type _} t.indexEquiv <| FreeAbelianGroup.basis _ #align free_abelian_group.equiv.of_free_abelian_group_linear_equiv FreeAbelianGroup.Equiv.ofFreeAbelianGroupLinearEquiv +/- warning: free_abelian_group.equiv.of_free_abelian_group_equiv -> FreeAbelianGroup.Equiv.ofFreeAbelianGroupEquiv is a dubious translation: +lean 3 declaration is + forall {α : Type.{u_1}} {β : Type.{u_2}}, (AddEquiv.{u_1, u_2} (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.{u_2} β) (AddZeroClass.toHasAdd.{u_1} (FreeAbelianGroup.{u_1} α) (AddMonoid.toAddZeroClass.{u_1} (FreeAbelianGroup.{u_1} α) (SubNegMonoid.toAddMonoid.{u_1} (FreeAbelianGroup.{u_1} α) (AddGroup.toSubNegMonoid.{u_1} (FreeAbelianGroup.{u_1} α) (AddCommGroup.toAddGroup.{u_1} (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.addCommGroup.{u_1} α)))))) (AddZeroClass.toHasAdd.{u_2} (FreeAbelianGroup.{u_2} β) (AddMonoid.toAddZeroClass.{u_2} (FreeAbelianGroup.{u_2} β) (SubNegMonoid.toAddMonoid.{u_2} (FreeAbelianGroup.{u_2} β) (AddGroup.toSubNegMonoid.{u_2} (FreeAbelianGroup.{u_2} β) (AddCommGroup.toAddGroup.{u_2} (FreeAbelianGroup.{u_2} β) (FreeAbelianGroup.addCommGroup.{u_2} β))))))) -> (Equiv.{succ u_1, succ u_2} α β) +but is expected to have type + forall {α : Type.{u_1}} {β : Type.{u_2}}, (AddEquiv.{u_1, u_2} (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.{u_2} β) (AddZeroClass.toAdd.{u_1} (FreeAbelianGroup.{u_1} α) (AddMonoid.toAddZeroClass.{u_1} (FreeAbelianGroup.{u_1} α) (SubNegMonoid.toAddMonoid.{u_1} (FreeAbelianGroup.{u_1} α) (AddGroup.toSubNegMonoid.{u_1} (FreeAbelianGroup.{u_1} α) (AddCommGroup.toAddGroup.{u_1} (FreeAbelianGroup.{u_1} α) (FreeAbelianGroup.addCommGroup.{u_1} α)))))) (AddZeroClass.toAdd.{u_2} (FreeAbelianGroup.{u_2} β) (AddMonoid.toAddZeroClass.{u_2} (FreeAbelianGroup.{u_2} β) (SubNegMonoid.toAddMonoid.{u_2} (FreeAbelianGroup.{u_2} β) (AddGroup.toSubNegMonoid.{u_2} (FreeAbelianGroup.{u_2} β) (AddCommGroup.toAddGroup.{u_2} (FreeAbelianGroup.{u_2} β) (FreeAbelianGroup.addCommGroup.{u_2} β))))))) -> (Equiv.{succ u_1, succ u_2} α β) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.equiv.of_free_abelian_group_equiv FreeAbelianGroup.Equiv.ofFreeAbelianGroupEquivₓ'. -/ /-- Isomorphic free abelian groups (as additive groups) have equivalent bases. -/ def Equiv.ofFreeAbelianGroupEquiv {α β : Type _} (e : FreeAbelianGroup α ≃+ FreeAbelianGroup β) : α ≃ β := Equiv.ofFreeAbelianGroupLinearEquiv e.toIntLinearEquiv #align free_abelian_group.equiv.of_free_abelian_group_equiv FreeAbelianGroup.Equiv.ofFreeAbelianGroupEquiv +/- warning: free_abelian_group.equiv.of_free_group_equiv -> FreeAbelianGroup.Equiv.ofFreeGroupEquiv is a dubious translation: +lean 3 declaration is + forall {α : Type.{u_1}} {β : Type.{u_2}}, (MulEquiv.{u_1, u_2} (FreeGroup.{u_1} α) (FreeGroup.{u_2} β) (FreeGroup.hasMul.{u_1} α) (FreeGroup.hasMul.{u_2} β)) -> (Equiv.{succ u_1, succ u_2} α β) +but is expected to have type + forall {α : Type.{u_1}} {β : Type.{u_2}}, (MulEquiv.{u_1, u_2} (FreeGroup.{u_1} α) (FreeGroup.{u_2} β) (FreeGroup.instMulFreeGroup.{u_1} α) (FreeGroup.instMulFreeGroup.{u_2} β)) -> (Equiv.{succ u_1, succ u_2} α β) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.equiv.of_free_group_equiv FreeAbelianGroup.Equiv.ofFreeGroupEquivₓ'. -/ /-- Isomorphic free groups have equivalent bases. -/ def Equiv.ofFreeGroupEquiv {α β : Type _} (e : FreeGroup α ≃* FreeGroup β) : α ≃ β := Equiv.ofFreeAbelianGroupEquiv e.abelianizationCongr.toAdditive @@ -138,6 +212,12 @@ def Equiv.ofFreeGroupEquiv {α β : Type _} (e : FreeGroup α ≃* FreeGroup β) open IsFreeGroup +/- warning: free_abelian_group.equiv.of_is_free_group_equiv -> FreeAbelianGroup.Equiv.ofIsFreeGroupEquiv is a dubious translation: +lean 3 declaration is + forall {G : Type.{u_1}} {H : Type.{u_2}} [_inst_1 : Group.{u_1} G] [_inst_2 : Group.{u_2} H] [_inst_3 : IsFreeGroup.{u_1} G _inst_1] [_inst_4 : IsFreeGroup.{u_2} H _inst_2], (MulEquiv.{u_1, u_2} G H (MulOneClass.toHasMul.{u_1} G (Monoid.toMulOneClass.{u_1} G (DivInvMonoid.toMonoid.{u_1} G (Group.toDivInvMonoid.{u_1} G _inst_1)))) (MulOneClass.toHasMul.{u_2} H (Monoid.toMulOneClass.{u_2} H (DivInvMonoid.toMonoid.{u_2} H (Group.toDivInvMonoid.{u_2} H _inst_2))))) -> (Equiv.{succ u_1, succ u_2} (IsFreeGroup.Generators.{u_1} G _inst_1 _inst_3) (IsFreeGroup.Generators.{u_2} H _inst_2 _inst_4)) +but is expected to have type + forall {G : Type.{u_1}} {H : Type.{u_2}} [_inst_1 : Group.{u_1} G] [_inst_2 : Group.{u_2} H] [_inst_3 : IsFreeGroup.{u_1} G _inst_1] [_inst_4 : IsFreeGroup.{u_2} H _inst_2], (MulEquiv.{u_1, u_2} G H (MulOneClass.toMul.{u_1} G (Monoid.toMulOneClass.{u_1} G (DivInvMonoid.toMonoid.{u_1} G (Group.toDivInvMonoid.{u_1} G _inst_1)))) (MulOneClass.toMul.{u_2} H (Monoid.toMulOneClass.{u_2} H (DivInvMonoid.toMonoid.{u_2} H (Group.toDivInvMonoid.{u_2} H _inst_2))))) -> (Equiv.{succ u_1, succ u_2} (IsFreeGroup.Generators.{u_1} G _inst_1 _inst_3) (IsFreeGroup.Generators.{u_2} H _inst_2 _inst_4)) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.equiv.of_is_free_group_equiv FreeAbelianGroup.Equiv.ofIsFreeGroupEquivₓ'. -/ /-- Isomorphic free groups have equivalent bases (`is_free_group` variant`). -/ def Equiv.ofIsFreeGroupEquiv {G H : Type _} [Group G] [Group H] [IsFreeGroup G] [IsFreeGroup H] (e : G ≃* H) : Generators G ≃ Generators H := @@ -146,45 +226,80 @@ def Equiv.ofIsFreeGroupEquiv {G H : Type _} [Group G] [Group H] [IsFreeGroup G] variable {X} +/- warning: free_abelian_group.coeff -> FreeAbelianGroup.coeff is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}}, X -> (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) +but is expected to have type + forall {X : Type.{u1}}, X -> (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.coeff FreeAbelianGroup.coeffₓ'. -/ /-- `coeff x` is the additive group homomorphism `free_abelian_group X →+ ℤ` that sends `a` to the multiplicity of `x : X` in `a`. -/ def coeff (x : X) : FreeAbelianGroup X →+ ℤ := (Finsupp.applyAddHom x).comp toFinsupp #align free_abelian_group.coeff FreeAbelianGroup.coeff +#print FreeAbelianGroup.support /- /-- `support a` for `a : free_abelian_group X` is the finite set of `x : X` that occur in the formal sum `a`. -/ def support (a : FreeAbelianGroup X) : Finset X := a.toFinsupp.support #align free_abelian_group.support FreeAbelianGroup.support +-/ +/- warning: free_abelian_group.mem_support_iff -> FreeAbelianGroup.mem_support_iff is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}} (x : X) (a : FreeAbelianGroup.{u1} X), Iff (Membership.Mem.{u1, u1} X (Finset.{u1} X) (Finset.hasMem.{u1} X) x (FreeAbelianGroup.support.{u1} X a)) (Ne.{1} Int (coeFn.{succ u1, succ u1} (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (fun (_x : AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) => (FreeAbelianGroup.{u1} X) -> Int) (AddMonoidHom.hasCoeToFun.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (FreeAbelianGroup.coeff.{u1} X x) a) (OfNat.ofNat.{0} Int 0 (OfNat.mk.{0} Int 0 (Zero.zero.{0} Int Int.hasZero)))) +but is expected to have type + forall {X : Type.{u1}} (x : X) (a : FreeAbelianGroup.{u1} X), Iff (Membership.mem.{u1, u1} X (Finset.{u1} X) (Finset.instMembershipFinset.{u1} X) x (FreeAbelianGroup.support.{u1} X a)) (Ne.{1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Int) a) (FunLike.coe.{succ u1, succ u1, 1} (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (FreeAbelianGroup.{u1} X) (fun (_x : FreeAbelianGroup.{u1} X) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Int) _x) (AddHomClass.toFunLike.{u1, u1, 0} (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (FreeAbelianGroup.{u1} X) Int (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddZeroClass.toAdd.{0} Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoidHomClass.toAddHomClass.{u1, u1, 0} (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt) (AddMonoidHom.addMonoidHomClass.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)))) (FreeAbelianGroup.coeff.{u1} X x) a) (OfNat.ofNat.{0} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Int) a) 0 (instOfNatInt 0))) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.mem_support_iff FreeAbelianGroup.mem_support_iffₓ'. -/ theorem mem_support_iff (x : X) (a : FreeAbelianGroup X) : x ∈ a.support ↔ coeff x a ≠ 0 := by rw [support, Finsupp.mem_support_iff] exact Iff.rfl #align free_abelian_group.mem_support_iff FreeAbelianGroup.mem_support_iff +/- warning: free_abelian_group.not_mem_support_iff -> FreeAbelianGroup.not_mem_support_iff is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}} (x : X) (a : FreeAbelianGroup.{u1} X), Iff (Not (Membership.Mem.{u1, u1} X (Finset.{u1} X) (Finset.hasMem.{u1} X) x (FreeAbelianGroup.support.{u1} X a))) (Eq.{1} Int (coeFn.{succ u1, succ u1} (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (fun (_x : AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) => (FreeAbelianGroup.{u1} X) -> Int) (AddMonoidHom.hasCoeToFun.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.addMonoid)) (FreeAbelianGroup.coeff.{u1} X x) a) (OfNat.ofNat.{0} Int 0 (OfNat.mk.{0} Int 0 (Zero.zero.{0} Int Int.hasZero)))) +but is expected to have type + forall {X : Type.{u1}} (x : X) (a : FreeAbelianGroup.{u1} X), Iff (Not (Membership.mem.{u1, u1} X (Finset.{u1} X) (Finset.instMembershipFinset.{u1} X) x (FreeAbelianGroup.support.{u1} X a))) (Eq.{1} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Int) a) (FunLike.coe.{succ u1, succ u1, 1} (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (FreeAbelianGroup.{u1} X) (fun (_x : FreeAbelianGroup.{u1} X) => (fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Int) _x) (AddHomClass.toFunLike.{u1, u1, 0} (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (FreeAbelianGroup.{u1} X) Int (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) (AddZeroClass.toAdd.{0} Int (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (AddMonoidHomClass.toAddHomClass.{u1, u1, 0} (AddMonoidHom.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)) (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt) (AddMonoidHom.addMonoidHomClass.{u1, 0} (FreeAbelianGroup.{u1} X) Int (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))) (AddMonoid.toAddZeroClass.{0} Int Int.instAddMonoidInt)))) (FreeAbelianGroup.coeff.{u1} X x) a) (OfNat.ofNat.{0} ((fun (x._@.Mathlib.Algebra.Hom.Group._hyg.403 : FreeAbelianGroup.{u1} X) => Int) a) 0 (instOfNatInt 0))) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.not_mem_support_iff FreeAbelianGroup.not_mem_support_iffₓ'. -/ theorem not_mem_support_iff (x : X) (a : FreeAbelianGroup X) : x ∉ a.support ↔ coeff x a = 0 := by rw [support, Finsupp.not_mem_support_iff] exact Iff.rfl #align free_abelian_group.not_mem_support_iff FreeAbelianGroup.not_mem_support_iff +/- warning: free_abelian_group.support_zero -> FreeAbelianGroup.support_zero is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}}, Eq.{succ u1} (Finset.{u1} X) (FreeAbelianGroup.support.{u1} X (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} X) 0 (OfNat.mk.{u1} (FreeAbelianGroup.{u1} X) 0 (Zero.zero.{u1} (FreeAbelianGroup.{u1} X) (AddZeroClass.toHasZero.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))))))) (EmptyCollection.emptyCollection.{u1} (Finset.{u1} X) (Finset.hasEmptyc.{u1} X)) +but is expected to have type + forall {X : Type.{u1}}, Eq.{succ u1} (Finset.{u1} X) (FreeAbelianGroup.support.{u1} X (OfNat.ofNat.{u1} (FreeAbelianGroup.{u1} X) 0 (Zero.toOfNat0.{u1} (FreeAbelianGroup.{u1} X) (NegZeroClass.toZero.{u1} (FreeAbelianGroup.{u1} X) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} X) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))))))) (EmptyCollection.emptyCollection.{u1} (Finset.{u1} X) (Finset.instEmptyCollectionFinset.{u1} X)) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.support_zero FreeAbelianGroup.support_zeroₓ'. -/ @[simp] theorem support_zero : support (0 : FreeAbelianGroup X) = ∅ := by simp only [support, Finsupp.support_zero, AddMonoidHom.map_zero] #align free_abelian_group.support_zero FreeAbelianGroup.support_zero +#print FreeAbelianGroup.support_of /- @[simp] theorem support_of (x : X) : support (of x) = {x} := by simp only [support, to_finsupp_of, Finsupp.support_single_ne_zero _ one_ne_zero] #align free_abelian_group.support_of FreeAbelianGroup.support_of +-/ +/- warning: free_abelian_group.support_neg -> FreeAbelianGroup.support_neg is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}} (a : FreeAbelianGroup.{u1} X), Eq.{succ u1} (Finset.{u1} X) (FreeAbelianGroup.support.{u1} X (Neg.neg.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toHasNeg.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))) a)) (FreeAbelianGroup.support.{u1} X a) +but is expected to have type + forall {X : Type.{u1}} (a : FreeAbelianGroup.{u1} X), Eq.{succ u1} (Finset.{u1} X) (FreeAbelianGroup.support.{u1} X (Neg.neg.{u1} (FreeAbelianGroup.{u1} X) (NegZeroClass.toNeg.{u1} (FreeAbelianGroup.{u1} X) (SubNegZeroMonoid.toNegZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubtractionMonoid.toSubNegZeroMonoid.{u1} (FreeAbelianGroup.{u1} X) (SubtractionCommMonoid.toSubtractionMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toDivisionAddCommMonoid.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X)))))) a)) (FreeAbelianGroup.support.{u1} X a) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.support_neg FreeAbelianGroup.support_negₓ'. -/ @[simp] theorem support_neg (a : FreeAbelianGroup X) : support (-a) = support a := by simp only [support, AddMonoidHom.map_neg, Finsupp.support_neg] #align free_abelian_group.support_neg FreeAbelianGroup.support_neg +#print FreeAbelianGroup.support_zsmul /- @[simp] theorem support_zsmul (k : ℤ) (h : k ≠ 0) (a : FreeAbelianGroup X) : support (k • a) = support a := by @@ -192,16 +307,25 @@ theorem support_zsmul (k : ℤ) (h : k ≠ 0) (a : FreeAbelianGroup X) : support simp only [mem_support_iff, AddMonoidHom.map_zsmul] simp only [h, zsmul_int_int, false_or_iff, Ne.def, mul_eq_zero] #align free_abelian_group.support_zsmul FreeAbelianGroup.support_zsmul +-/ +#print FreeAbelianGroup.support_nsmul /- @[simp] theorem support_nsmul (k : ℕ) (h : k ≠ 0) (a : FreeAbelianGroup X) : support (k • a) = support a := by apply support_zsmul k _ a exact_mod_cast h #align free_abelian_group.support_nsmul FreeAbelianGroup.support_nsmul +-/ open Classical +/- warning: free_abelian_group.support_add -> FreeAbelianGroup.support_add is a dubious translation: +lean 3 declaration is + forall {X : Type.{u1}} (a : FreeAbelianGroup.{u1} X) (b : FreeAbelianGroup.{u1} X), HasSubset.Subset.{u1} (Finset.{u1} X) (Finset.hasSubset.{u1} X) (FreeAbelianGroup.support.{u1} X (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.{u1} X) (instHAdd.{u1} (FreeAbelianGroup.{u1} X) (AddZeroClass.toHasAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))))) a b)) (Union.union.{u1} (Finset.{u1} X) (Finset.hasUnion.{u1} X (fun (a : X) (b : X) => Classical.propDecidable (Eq.{succ u1} X a b))) (FreeAbelianGroup.support.{u1} X a) (FreeAbelianGroup.support.{u1} X b)) +but is expected to have type + forall {X : Type.{u1}} (a : FreeAbelianGroup.{u1} X) (b : FreeAbelianGroup.{u1} X), HasSubset.Subset.{u1} (Finset.{u1} X) (Finset.instHasSubsetFinset.{u1} X) (FreeAbelianGroup.support.{u1} X (HAdd.hAdd.{u1, u1, u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.{u1} X) (instHAdd.{u1} (FreeAbelianGroup.{u1} X) (AddZeroClass.toAdd.{u1} (FreeAbelianGroup.{u1} X) (AddMonoid.toAddZeroClass.{u1} (FreeAbelianGroup.{u1} X) (SubNegMonoid.toAddMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddGroup.toSubNegMonoid.{u1} (FreeAbelianGroup.{u1} X) (AddCommGroup.toAddGroup.{u1} (FreeAbelianGroup.{u1} X) (FreeAbelianGroup.addCommGroup.{u1} X))))))) a b)) (Union.union.{u1} (Finset.{u1} X) (Finset.instUnionFinset.{u1} X (fun (a : X) (b : X) => Classical.propDecidable (Eq.{succ u1} X a b))) (FreeAbelianGroup.support.{u1} X a) (FreeAbelianGroup.support.{u1} X b)) +Case conversion may be inaccurate. Consider using '#align free_abelian_group.support_add FreeAbelianGroup.support_addₓ'. -/ theorem support_add (a b : FreeAbelianGroup X) : support (a + b) ⊆ a.support ∪ b.support := by simp only [support, AddMonoidHom.map_add] diff --git a/Mathbin/GroupTheory/OrderOfElement.lean b/Mathbin/GroupTheory/OrderOfElement.lean index 003fa6f76f..b88ec64782 100644 --- a/Mathbin/GroupTheory/OrderOfElement.lean +++ b/Mathbin/GroupTheory/OrderOfElement.lean @@ -58,7 +58,7 @@ section IsOfFinOrder lean 3 declaration is forall {G : Type.{u1}} {n : Nat} [_inst_1 : Monoid.{u1} G] (x : G), Iff (Function.IsPeriodicPt.{u1} G (HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1))) x) n (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1)))))) (Eq.{succ u1} G (HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G _inst_1)) x n) (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1)))))) but is expected to have type - forall {G : Type.{u1}} {n : Nat} [_inst_1 : Monoid.{u1} G] (x : G), Iff (Function.IsPeriodicPt.{u1} G ((fun (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.83 : G) (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.85 : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1))) x._@.Mathlib.GroupTheory.OrderOfElement._hyg.83 x._@.Mathlib.GroupTheory.OrderOfElement._hyg.85) x) n (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (Monoid.toOne.{u1} G _inst_1)))) (Eq.{succ u1} G (HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G _inst_1)) x n) (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (Monoid.toOne.{u1} G _inst_1)))) + forall {G : Type.{u1}} {n : Nat} [_inst_1 : Monoid.{u1} G] (x : G), Iff (Function.IsPeriodicPt.{u1} G ((fun (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.82 : G) (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.84 : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1))) x._@.Mathlib.GroupTheory.OrderOfElement._hyg.82 x._@.Mathlib.GroupTheory.OrderOfElement._hyg.84) x) n (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (Monoid.toOne.{u1} G _inst_1)))) (Eq.{succ u1} G (HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G _inst_1)) x n) (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (Monoid.toOne.{u1} G _inst_1)))) Case conversion may be inaccurate. Consider using '#align is_periodic_pt_mul_iff_pow_eq_one isPeriodicPt_mul_iff_pow_eq_oneₓ'. -/ @[to_additive] theorem isPeriodicPt_mul_iff_pow_eq_one (x : G) : IsPeriodicPt ((· * ·) x) n 1 ↔ x ^ n = 1 := by @@ -718,7 +718,7 @@ theorem pow_injective_of_lt_orderOf (hn : n < orderOf x) (hm : m < orderOf x) (e lean 3 declaration is forall {G : Type.{u1}} (x : G) (y : G) [_inst_1 : LeftCancelMonoid.{u1} G] [_inst_2 : DecidableEq.{succ u1} G], (LT.lt.{0} Nat Nat.hasLt (OfNat.ofNat.{0} Nat 0 (OfNat.mk.{0} Nat 0 (Zero.zero.{0} Nat Nat.hasZero))) (orderOf.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1) x)) -> (Iff (Membership.Mem.{u1, u1} G (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) (SetLike.hasMem.{u1, u1} (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) G (Submonoid.setLike.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1)))) y (Submonoid.powers.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1) x)) (Membership.Mem.{u1, u1} G (Finset.{u1} G) (Finset.hasMem.{u1} G) y (Finset.image.{0, u1} Nat G (fun (a : G) (b : G) => _inst_2 a b) (HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) x) (Finset.range (orderOf.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1) x))))) but is expected to have type - forall {G : Type.{u1}} {x : G} {y : G} [_inst_1 : Monoid.{u1} G] [_inst_2 : DecidableEq.{succ u1} G], (LT.lt.{0} Nat instLTNat (OfNat.ofNat.{0} Nat 0 (instOfNatNat 0)) (orderOf.{u1} G _inst_1 x)) -> (Iff (Membership.mem.{u1, u1} G (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1)) (SetLike.instMembership.{u1, u1} (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1)) G (Submonoid.instSetLikeSubmonoid.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1))) y (Submonoid.powers.{u1} G _inst_1 x)) (Membership.mem.{u1, u1} G (Finset.{u1} G) (Finset.instMembershipFinset.{u1} G) y (Finset.image.{0, u1} Nat G (fun (a : G) (b : G) => _inst_2 a b) (fun (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.2063 : Nat) => HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G _inst_1)) x x._@.Mathlib.GroupTheory.OrderOfElement._hyg.2063) (Finset.range (orderOf.{u1} G _inst_1 x))))) + forall {G : Type.{u1}} {x : G} {y : G} [_inst_1 : Monoid.{u1} G] [_inst_2 : DecidableEq.{succ u1} G], (LT.lt.{0} Nat instLTNat (OfNat.ofNat.{0} Nat 0 (instOfNatNat 0)) (orderOf.{u1} G _inst_1 x)) -> (Iff (Membership.mem.{u1, u1} G (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1)) (SetLike.instMembership.{u1, u1} (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1)) G (Submonoid.instSetLikeSubmonoid.{u1} G (Monoid.toMulOneClass.{u1} G _inst_1))) y (Submonoid.powers.{u1} G _inst_1 x)) (Membership.mem.{u1, u1} G (Finset.{u1} G) (Finset.instMembershipFinset.{u1} G) y (Finset.image.{0, u1} Nat G (fun (a : G) (b : G) => _inst_2 a b) (fun (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.2062 : Nat) => HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G _inst_1)) x x._@.Mathlib.GroupTheory.OrderOfElement._hyg.2062) (Finset.range (orderOf.{u1} G _inst_1 x))))) Case conversion may be inaccurate. Consider using '#align mem_powers_iff_mem_range_order_of' mem_powers_iff_mem_range_order_of'ₓ'. -/ @[to_additive mem_multiples_iff_mem_range_addOrderOf'] theorem mem_powers_iff_mem_range_order_of' [DecidableEq G] (hx : 0 < orderOf x) : @@ -1111,7 +1111,7 @@ theorem orderOf_pow [Finite G] (x : G) : orderOf (x ^ n) = orderOf x / gcd (orde lean 3 declaration is forall {G : Type.{u1}} {x : G} {y : G} [_inst_1 : LeftCancelMonoid.{u1} G] [_inst_3 : Finite.{succ u1} G] [_inst_4 : DecidableEq.{succ u1} G], Iff (Membership.Mem.{u1, u1} G (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) (SetLike.hasMem.{u1, u1} (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) G (Submonoid.setLike.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1)))) y (Submonoid.powers.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1) x)) (Membership.Mem.{u1, u1} G (Finset.{u1} G) (Finset.hasMem.{u1} G) y (Finset.image.{0, u1} Nat G (fun (a : G) (b : G) => _inst_4 a b) (HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) x) (Finset.range (orderOf.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1) x)))) but is expected to have type - forall {G : Type.{u1}} {x : G} {y : G} [_inst_1 : LeftCancelMonoid.{u1} G] [_inst_3 : Finite.{succ u1} G] [_inst_4 : DecidableEq.{succ u1} G], Iff (Membership.mem.{u1, u1} G (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) (SetLike.instMembership.{u1, u1} (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) G (Submonoid.instSetLikeSubmonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1)))) y (Submonoid.powers.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1) x)) (Membership.mem.{u1, u1} G (Finset.{u1} G) (Finset.instMembershipFinset.{u1} G) y (Finset.image.{0, u1} Nat G (fun (a : G) (b : G) => _inst_4 a b) ((fun (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.6916 : G) (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.6918 : Nat) => HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) x._@.Mathlib.GroupTheory.OrderOfElement._hyg.6916 x._@.Mathlib.GroupTheory.OrderOfElement._hyg.6918) x) (Finset.range (orderOf.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1) x)))) + forall {G : Type.{u1}} {x : G} {y : G} [_inst_1 : LeftCancelMonoid.{u1} G] [_inst_3 : Finite.{succ u1} G] [_inst_4 : DecidableEq.{succ u1} G], Iff (Membership.mem.{u1, u1} G (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) (SetLike.instMembership.{u1, u1} (Submonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) G (Submonoid.instSetLikeSubmonoid.{u1} G (Monoid.toMulOneClass.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1)))) y (Submonoid.powers.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1) x)) (Membership.mem.{u1, u1} G (Finset.{u1} G) (Finset.instMembershipFinset.{u1} G) y (Finset.image.{0, u1} Nat G (fun (a : G) (b : G) => _inst_4 a b) ((fun (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.6915 : G) (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.6917 : Nat) => HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1))) x._@.Mathlib.GroupTheory.OrderOfElement._hyg.6915 x._@.Mathlib.GroupTheory.OrderOfElement._hyg.6917) x) (Finset.range (orderOf.{u1} G (LeftCancelMonoid.toMonoid.{u1} G _inst_1) x)))) Case conversion may be inaccurate. Consider using '#align mem_powers_iff_mem_range_order_of mem_powers_iff_mem_range_orderOfₓ'. -/ @[to_additive mem_multiples_iff_mem_range_addOrderOf] theorem mem_powers_iff_mem_range_orderOf [Finite G] [DecidableEq G] : @@ -1273,7 +1273,7 @@ theorem powers_eq_zpowers [Finite G] (x : G) : (Submonoid.powers x : Set G) = zp lean 3 declaration is forall {G : Type.{u1}} {x : G} {y : G} [_inst_1 : Group.{u1} G] [_inst_3 : Finite.{succ u1} G] [_inst_4 : DecidableEq.{succ u1} G], Iff (Membership.Mem.{u1, u1} G (Subgroup.{u1} G _inst_1) (SetLike.hasMem.{u1, u1} (Subgroup.{u1} G _inst_1) G (Subgroup.setLike.{u1} G _inst_1)) y (Subgroup.zpowers.{u1} G _inst_1 x)) (Membership.Mem.{u1, u1} G (Finset.{u1} G) (Finset.hasMem.{u1} G) y (Finset.image.{0, u1} Nat G (fun (a : G) (b : G) => _inst_4 a b) (HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1)))) x) (Finset.range (orderOf.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1)) x)))) but is expected to have type - forall {G : Type.{u1}} {x : G} {y : G} [_inst_1 : Group.{u1} G] [_inst_3 : Finite.{succ u1} G] [_inst_4 : DecidableEq.{succ u1} G], Iff (Membership.mem.{u1, u1} G (Subgroup.{u1} G _inst_1) (SetLike.instMembership.{u1, u1} (Subgroup.{u1} G _inst_1) G (Subgroup.instSetLikeSubgroup.{u1} G _inst_1)) y (Subgroup.zpowers.{u1} G _inst_1 x)) (Membership.mem.{u1, u1} G (Finset.{u1} G) (Finset.instMembershipFinset.{u1} G) y (Finset.image.{0, u1} Nat G (fun (a : G) (b : G) => _inst_4 a b) ((fun (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.7973 : G) (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.7975 : Nat) => HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1)))) x._@.Mathlib.GroupTheory.OrderOfElement._hyg.7973 x._@.Mathlib.GroupTheory.OrderOfElement._hyg.7975) x) (Finset.range (orderOf.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1)) x)))) + forall {G : Type.{u1}} {x : G} {y : G} [_inst_1 : Group.{u1} G] [_inst_3 : Finite.{succ u1} G] [_inst_4 : DecidableEq.{succ u1} G], Iff (Membership.mem.{u1, u1} G (Subgroup.{u1} G _inst_1) (SetLike.instMembership.{u1, u1} (Subgroup.{u1} G _inst_1) G (Subgroup.instSetLikeSubgroup.{u1} G _inst_1)) y (Subgroup.zpowers.{u1} G _inst_1 x)) (Membership.mem.{u1, u1} G (Finset.{u1} G) (Finset.instMembershipFinset.{u1} G) y (Finset.image.{0, u1} Nat G (fun (a : G) (b : G) => _inst_4 a b) ((fun (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.7972 : G) (x._@.Mathlib.GroupTheory.OrderOfElement._hyg.7974 : Nat) => HPow.hPow.{u1, 0, u1} G Nat G (instHPow.{u1, 0} G Nat (Monoid.Pow.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1)))) x._@.Mathlib.GroupTheory.OrderOfElement._hyg.7972 x._@.Mathlib.GroupTheory.OrderOfElement._hyg.7974) x) (Finset.range (orderOf.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1)) x)))) Case conversion may be inaccurate. Consider using '#align mem_zpowers_iff_mem_range_order_of mem_zpowers_iff_mem_range_orderOfₓ'. -/ @[to_additive mem_zmultiples_iff_mem_range_addOrderOf] theorem mem_zpowers_iff_mem_range_orderOf [Finite G] [DecidableEq G] : diff --git a/Mathbin/Order/Category/BddDistLat.lean b/Mathbin/Order/Category/BddDistLat.lean index f595698702..3b37d98272 100644 --- a/Mathbin/Order/Category/BddDistLat.lean +++ b/Mathbin/Order/Category/BddDistLat.lean @@ -27,7 +27,7 @@ open CategoryTheory /-- The category of bounded distributive lattices with bounded lattice morphisms. -/ structure BddDistLat where - toDistLat : DistLat + toDistLat : DistLatCat [isBoundedOrder : BoundedOrder to_DistLat] #align BddDistLat BddDistLat @@ -70,7 +70,7 @@ instance : LargeCategory.{u} BddDistLat := instance : ConcreteCategory BddDistLat := InducedCategory.concreteCategory toBddLat -instance hasForgetToDistLat : HasForget₂ BddDistLat DistLat +instance hasForgetToDistLat : HasForget₂ BddDistLat DistLatCat where forget₂ := { obj := fun X => ⟨X⟩ map := fun X Y => BoundedLatticeHom.toLatticeHom } @@ -80,11 +80,11 @@ instance hasForgetToBddLat : HasForget₂ BddDistLat BddLat := InducedCategory.hasForget₂ toBddLat #align BddDistLat.has_forget_to_BddLat BddDistLat.hasForgetToBddLat -theorem forget_bddLat_latCat_eq_forget_distLat_latCat : +theorem forget_bddLat_latCat_eq_forget_distLatCat_latCat : forget₂ BddDistLat BddLat ⋙ forget₂ BddLat LatCat = - forget₂ BddDistLat DistLat ⋙ forget₂ DistLat LatCat := + forget₂ BddDistLat DistLatCat ⋙ forget₂ DistLatCat LatCat := rfl -#align BddDistLat.forget_BddLat_Lat_eq_forget_DistLat_Lat BddDistLat.forget_bddLat_latCat_eq_forget_distLat_latCat +#align BddDistLat.forget_BddLat_Lat_eq_forget_DistLat_Lat BddDistLat.forget_bddLat_latCat_eq_forget_distLatCat_latCat /-- Constructs an equivalence between bounded distributive lattices from an order isomorphism between them. -/ @@ -118,8 +118,9 @@ def dualEquiv : BddDistLat ≌ BddDistLat := end BddDistLat -theorem bddDistLat_dual_comp_forget_to_distLat : - BddDistLat.dual ⋙ forget₂ BddDistLat DistLat = forget₂ BddDistLat DistLat ⋙ DistLat.dual := +theorem bddDistLat_dual_comp_forget_to_distLatCat : + BddDistLat.dual ⋙ forget₂ BddDistLat DistLatCat = + forget₂ BddDistLat DistLatCat ⋙ DistLatCat.dual := rfl -#align BddDistLat_dual_comp_forget_to_DistLat bddDistLat_dual_comp_forget_to_distLat +#align BddDistLat_dual_comp_forget_to_DistLat bddDistLat_dual_comp_forget_to_distLatCat diff --git a/Mathbin/Order/Category/DistLat.lean b/Mathbin/Order/Category/DistLat.lean index a849d535eb..aa67d7f034 100644 --- a/Mathbin/Order/Category/DistLat.lean +++ b/Mathbin/Order/Category/DistLat.lean @@ -25,45 +25,63 @@ universe u open CategoryTheory +#print DistLatCat /- /-- The category of distributive lattices. -/ -def DistLat := +def DistLatCat := Bundled DistribLattice -#align DistLat DistLat +#align DistLat DistLatCat +-/ -namespace DistLat +namespace DistLatCat -instance : CoeSort DistLat (Type _) := +instance : CoeSort DistLatCat (Type _) := Bundled.hasCoeToSort -instance (X : DistLat) : DistribLattice X := +instance (X : DistLatCat) : DistribLattice X := X.str +#print DistLatCat.of /- /-- Construct a bundled `DistLat` from a `distrib_lattice` underlying type and typeclass. -/ -def of (α : Type _) [DistribLattice α] : DistLat := +def of (α : Type _) [DistribLattice α] : DistLatCat := Bundled.of α -#align DistLat.of DistLat.of +#align DistLat.of DistLatCat.of +-/ +#print DistLatCat.coe_of /- @[simp] theorem coe_of (α : Type _) [DistribLattice α] : ↥(of α) = α := rfl -#align DistLat.coe_of DistLat.coe_of +#align DistLat.coe_of DistLatCat.coe_of +-/ -instance : Inhabited DistLat := +instance : Inhabited DistLatCat := ⟨of PUnit⟩ instance : BundledHom.ParentProjection @DistribLattice.toLattice := ⟨⟩ -deriving instance LargeCategory, ConcreteCategory for DistLat +deriving instance LargeCategory, ConcreteCategory for DistLatCat -instance hasForgetToLat : HasForget₂ DistLat LatCat := +/- warning: DistLat.has_forget_to_Lat -> DistLatCat.hasForgetToLatCat is a dubious translation: +lean 3 declaration is + CategoryTheory.HasForget₂.{succ u1, succ u1, u1, u1, u1} DistLatCat.{u1} LatCat.{u1} DistLatCat.largeCategory.{u1} DistLatCat.concreteCategory.{u1} LatCat.CategoryTheory.largeCategory.{u1} LatCat.CategoryTheory.concreteCategory.{u1} +but is expected to have type + CategoryTheory.HasForget₂.{succ u1, succ u1, u1, u1, u1} DistLatCat.{u1} LatCat.{u1} instDistLatCatLargeCategory.{u1} DistLatCat.instConcreteCategoryDistLatCatInstDistLatCatLargeCategory.{u1} LatCat.instLargeCategoryLatCat.{u1} LatCat.instConcreteCategoryLatCatInstLargeCategoryLatCat.{u1} +Case conversion may be inaccurate. Consider using '#align DistLat.has_forget_to_Lat DistLatCat.hasForgetToLatCatₓ'. -/ +instance hasForgetToLatCat : HasForget₂ DistLatCat LatCat := BundledHom.forget₂ _ _ -#align DistLat.has_forget_to_Lat DistLat.hasForgetToLat - +#align DistLat.has_forget_to_Lat DistLatCat.hasForgetToLatCat + +/- warning: DistLat.iso.mk -> DistLatCat.Iso.mk is a dubious translation: +lean 3 declaration is + forall {α : DistLatCat.{u1}} {β : DistLatCat.{u1}}, (OrderIso.{u1, u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} α) (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} β) (Preorder.toLE.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} α) (PartialOrder.toPreorder.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} α) (SemilatticeInf.toPartialOrder.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} α) (Lattice.toSemilatticeInf.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} α) (DistribLattice.toLattice.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} α) (DistLatCat.distribLattice.{u1} α)))))) (Preorder.toLE.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} β) (PartialOrder.toPreorder.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} β) (SemilatticeInf.toPartialOrder.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} β) (Lattice.toSemilatticeInf.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} β) (DistribLattice.toLattice.{u1} (coeSort.{succ (succ u1), succ (succ u1)} DistLatCat.{u1} Type.{u1} DistLatCat.hasCoeToSort.{u1} β) (DistLatCat.distribLattice.{u1} β))))))) -> (CategoryTheory.Iso.{u1, succ u1} DistLatCat.{u1} DistLatCat.largeCategory.{u1} α β) +but is expected to have type + forall {α : DistLatCat.{u1}} {β : DistLatCat.{u1}}, (OrderIso.{u1, u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} α) (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} β) (Preorder.toLE.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} α) (PartialOrder.toPreorder.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} α) (SemilatticeInf.toPartialOrder.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} α) (Lattice.toSemilatticeInf.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} α) (DistribLattice.toLattice.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} α) (DistLatCat.instDistribLatticeα.{u1} α)))))) (Preorder.toLE.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} β) (PartialOrder.toPreorder.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} β) (SemilatticeInf.toPartialOrder.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} β) (Lattice.toSemilatticeInf.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} β) (DistribLattice.toLattice.{u1} (CategoryTheory.Bundled.α.{u1, u1} DistribLattice.{u1} β) (DistLatCat.instDistribLatticeα.{u1} β))))))) -> (CategoryTheory.Iso.{u1, succ u1} DistLatCat.{u1} instDistLatCatLargeCategory.{u1} α β) +Case conversion may be inaccurate. Consider using '#align DistLat.iso.mk DistLatCat.Iso.mkₓ'. -/ /-- Constructs an equivalence between distributive lattices from an order isomorphism between them. -/ @[simps] -def Iso.mk {α β : DistLat.{u}} (e : α ≃o β) : α ≅ β +def Iso.mk {α β : DistLatCat.{u}} (e : α ≃o β) : α ≅ β where Hom := e inv := e.symm @@ -73,27 +91,45 @@ def Iso.mk {α β : DistLat.{u}} (e : α ≃o β) : α ≅ β inv_hom_id' := by ext exact e.apply_symm_apply _ -#align DistLat.iso.mk DistLat.Iso.mk - +#align DistLat.iso.mk DistLatCat.Iso.mk + +/- warning: DistLat.dual -> DistLatCat.dual is a dubious translation: +lean 3 declaration is + CategoryTheory.Functor.{u1, u1, succ u1, succ u1} DistLatCat.{u1} DistLatCat.largeCategory.{u1} DistLatCat.{u1} DistLatCat.largeCategory.{u1} +but is expected to have type + CategoryTheory.Functor.{u1, u1, succ u1, succ u1} DistLatCat.{u1} instDistLatCatLargeCategory.{u1} DistLatCat.{u1} instDistLatCatLargeCategory.{u1} +Case conversion may be inaccurate. Consider using '#align DistLat.dual DistLatCat.dualₓ'. -/ /-- `order_dual` as a functor. -/ @[simps] -def dual : DistLat ⥤ DistLat where +def dual : DistLatCat ⥤ DistLatCat where obj X := of Xᵒᵈ map X Y := LatticeHom.dual -#align DistLat.dual DistLat.dual - +#align DistLat.dual DistLatCat.dual + +/- warning: DistLat.dual_equiv -> DistLatCat.dualEquiv is a dubious translation: +lean 3 declaration is + CategoryTheory.Equivalence.{u1, u1, succ u1, succ u1} DistLatCat.{u1} DistLatCat.largeCategory.{u1} DistLatCat.{u1} DistLatCat.largeCategory.{u1} +but is expected to have type + CategoryTheory.Equivalence.{u1, u1, succ u1, succ u1} DistLatCat.{u1} DistLatCat.{u1} instDistLatCatLargeCategory.{u1} instDistLatCatLargeCategory.{u1} +Case conversion may be inaccurate. Consider using '#align DistLat.dual_equiv DistLatCat.dualEquivₓ'. -/ /-- The equivalence between `DistLat` and itself induced by `order_dual` both ways. -/ @[simps Functor inverse] -def dualEquiv : DistLat ≌ DistLat := +def dualEquiv : DistLatCat ≌ DistLatCat := Equivalence.mk dual dual (NatIso.ofComponents (fun X => Iso.mk <| OrderIso.dualDual X) fun X Y f => rfl) (NatIso.ofComponents (fun X => Iso.mk <| OrderIso.dualDual X) fun X Y f => rfl) -#align DistLat.dual_equiv DistLat.dualEquiv - -end DistLat - -theorem distLat_dual_comp_forget_to_latCat : - DistLat.dual ⋙ forget₂ DistLat LatCat = forget₂ DistLat LatCat ⋙ LatCat.dual := +#align DistLat.dual_equiv DistLatCat.dualEquiv + +end DistLatCat + +/- warning: DistLat_dual_comp_forget_to_Lat -> distLatCat_dual_comp_forget_to_latCat is a dubious translation: +lean 3 declaration is + Eq.{succ (succ u1)} (CategoryTheory.Functor.{u1, u1, succ u1, succ u1} DistLatCat.{u1} DistLatCat.largeCategory.{u1} LatCat.{u1} LatCat.CategoryTheory.largeCategory.{u1}) (CategoryTheory.Functor.comp.{u1, u1, u1, succ u1, succ u1, succ u1} DistLatCat.{u1} DistLatCat.largeCategory.{u1} DistLatCat.{u1} DistLatCat.largeCategory.{u1} LatCat.{u1} LatCat.CategoryTheory.largeCategory.{u1} DistLatCat.dual.{u1} (CategoryTheory.forget₂.{succ u1, succ u1, u1, u1, u1} DistLatCat.{u1} LatCat.{u1} DistLatCat.largeCategory.{u1} DistLatCat.concreteCategory.{u1} LatCat.CategoryTheory.largeCategory.{u1} LatCat.CategoryTheory.concreteCategory.{u1} DistLatCat.hasForgetToLatCat.{u1})) (CategoryTheory.Functor.comp.{u1, u1, u1, succ u1, succ u1, succ u1} DistLatCat.{u1} DistLatCat.largeCategory.{u1} LatCat.{u1} LatCat.CategoryTheory.largeCategory.{u1} LatCat.{u1} LatCat.CategoryTheory.largeCategory.{u1} (CategoryTheory.forget₂.{succ u1, succ u1, u1, u1, u1} DistLatCat.{u1} LatCat.{u1} DistLatCat.largeCategory.{u1} DistLatCat.concreteCategory.{u1} LatCat.CategoryTheory.largeCategory.{u1} LatCat.CategoryTheory.concreteCategory.{u1} DistLatCat.hasForgetToLatCat.{u1}) LatCat.dual.{u1}) +but is expected to have type + Eq.{succ (succ u1)} (CategoryTheory.Functor.{u1, u1, succ u1, succ u1} DistLatCat.{u1} instDistLatCatLargeCategory.{u1} LatCat.{u1} LatCat.instLargeCategoryLatCat.{u1}) (CategoryTheory.Functor.comp.{u1, u1, u1, succ u1, succ u1, succ u1} DistLatCat.{u1} instDistLatCatLargeCategory.{u1} DistLatCat.{u1} instDistLatCatLargeCategory.{u1} LatCat.{u1} LatCat.instLargeCategoryLatCat.{u1} DistLatCat.dual.{u1} (CategoryTheory.forget₂.{succ u1, succ u1, u1, u1, u1} DistLatCat.{u1} LatCat.{u1} instDistLatCatLargeCategory.{u1} DistLatCat.instConcreteCategoryDistLatCatInstDistLatCatLargeCategory.{u1} LatCat.instLargeCategoryLatCat.{u1} LatCat.instConcreteCategoryLatCatInstLargeCategoryLatCat.{u1} DistLatCat.hasForgetToLatCat.{u1})) (CategoryTheory.Functor.comp.{u1, u1, u1, succ u1, succ u1, succ u1} DistLatCat.{u1} instDistLatCatLargeCategory.{u1} LatCat.{u1} LatCat.instLargeCategoryLatCat.{u1} LatCat.{u1} LatCat.instLargeCategoryLatCat.{u1} (CategoryTheory.forget₂.{succ u1, succ u1, u1, u1, u1} DistLatCat.{u1} LatCat.{u1} instDistLatCatLargeCategory.{u1} DistLatCat.instConcreteCategoryDistLatCatInstDistLatCatLargeCategory.{u1} LatCat.instLargeCategoryLatCat.{u1} LatCat.instConcreteCategoryLatCatInstLargeCategoryLatCat.{u1} DistLatCat.hasForgetToLatCat.{u1}) LatCat.dual.{u1}) +Case conversion may be inaccurate. Consider using '#align DistLat_dual_comp_forget_to_Lat distLatCat_dual_comp_forget_to_latCatₓ'. -/ +theorem distLatCat_dual_comp_forget_to_latCat : + DistLatCat.dual ⋙ forget₂ DistLatCat LatCat = forget₂ DistLatCat LatCat ⋙ LatCat.dual := rfl -#align DistLat_dual_comp_forget_to_Lat distLat_dual_comp_forget_to_latCat +#align DistLat_dual_comp_forget_to_Lat distLatCat_dual_comp_forget_to_latCat diff --git a/Mathbin/Order/Filter/Pointwise.lean b/Mathbin/Order/Filter/Pointwise.lean index 28ccf50117..119c97dea4 100644 --- a/Mathbin/Order/Filter/Pointwise.lean +++ b/Mathbin/Order/Filter/Pointwise.lean @@ -518,7 +518,7 @@ theorem le_mul_iff : h ≤ f * g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ lean 3 declaration is forall {α : Type.{u1}} [_inst_1 : Mul.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (HMul.hMul.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHMul.{u1} (Filter.{u1} α) (Filter.instMul.{u1} α _inst_1))) (LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.partialOrder.{u1} α)))) but is expected to have type - forall {α : Type.{u1}} [_inst_1 : Mul.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.2295 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.2297 : Filter.{u1} α) => HMul.hMul.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHMul.{u1} (Filter.{u1} α) (Filter.instMul.{u1} α _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.2295 x._@.Mathlib.Order.Filter.Pointwise._hyg.2297) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.2310 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.2312 : Filter.{u1} α) => LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.instPartialOrderFilter.{u1} α))) x._@.Mathlib.Order.Filter.Pointwise._hyg.2310 x._@.Mathlib.Order.Filter.Pointwise._hyg.2312) + forall {α : Type.{u1}} [_inst_1 : Mul.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.2294 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.2296 : Filter.{u1} α) => HMul.hMul.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHMul.{u1} (Filter.{u1} α) (Filter.instMul.{u1} α _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.2294 x._@.Mathlib.Order.Filter.Pointwise._hyg.2296) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.2309 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.2311 : Filter.{u1} α) => LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.instPartialOrderFilter.{u1} α))) x._@.Mathlib.Order.Filter.Pointwise._hyg.2309 x._@.Mathlib.Order.Filter.Pointwise._hyg.2311) Case conversion may be inaccurate. Consider using '#align filter.covariant_mul Filter.covariant_mulₓ'. -/ @[to_additive] instance covariant_mul : CovariantClass (Filter α) (Filter α) (· * ·) (· ≤ ·) := @@ -530,7 +530,7 @@ instance covariant_mul : CovariantClass (Filter α) (Filter α) (· * ·) (· lean 3 declaration is forall {α : Type.{u1}} [_inst_1 : Mul.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Function.swap.{succ u1, succ u1, succ u1} (Filter.{u1} α) (Filter.{u1} α) (fun (ᾰ : Filter.{u1} α) (ᾰ : Filter.{u1} α) => Filter.{u1} α) (HMul.hMul.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHMul.{u1} (Filter.{u1} α) (Filter.instMul.{u1} α _inst_1)))) (LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.partialOrder.{u1} α)))) but is expected to have type - forall {α : Type.{u1}} [_inst_1 : Mul.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Function.swap.{succ u1, succ u1, succ u1} (Filter.{u1} α) (Filter.{u1} α) (fun (ᾰ : Filter.{u1} α) (ᾰ : Filter.{u1} α) => Filter.{u1} α) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.2383 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.2385 : Filter.{u1} α) => HMul.hMul.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHMul.{u1} (Filter.{u1} α) (Filter.instMul.{u1} α _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.2383 x._@.Mathlib.Order.Filter.Pointwise._hyg.2385)) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.2398 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.2400 : Filter.{u1} α) => LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.instPartialOrderFilter.{u1} α))) x._@.Mathlib.Order.Filter.Pointwise._hyg.2398 x._@.Mathlib.Order.Filter.Pointwise._hyg.2400) + forall {α : Type.{u1}} [_inst_1 : Mul.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Function.swap.{succ u1, succ u1, succ u1} (Filter.{u1} α) (Filter.{u1} α) (fun (ᾰ : Filter.{u1} α) (ᾰ : Filter.{u1} α) => Filter.{u1} α) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.2382 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.2384 : Filter.{u1} α) => HMul.hMul.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHMul.{u1} (Filter.{u1} α) (Filter.instMul.{u1} α _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.2382 x._@.Mathlib.Order.Filter.Pointwise._hyg.2384)) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.2397 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.2399 : Filter.{u1} α) => LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.instPartialOrderFilter.{u1} α))) x._@.Mathlib.Order.Filter.Pointwise._hyg.2397 x._@.Mathlib.Order.Filter.Pointwise._hyg.2399) Case conversion may be inaccurate. Consider using '#align filter.covariant_swap_mul Filter.covariant_swap_mulₓ'. -/ @[to_additive] instance covariant_swap_mul : CovariantClass (Filter α) (Filter α) (swap (· * ·)) (· ≤ ·) := @@ -763,7 +763,7 @@ protected theorem le_div_iff : h ≤ f / g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t lean 3 declaration is forall {α : Type.{u1}} [_inst_1 : Div.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (HDiv.hDiv.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHDiv.{u1} (Filter.{u1} α) (Filter.instDiv.{u1} α _inst_1))) (LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.partialOrder.{u1} α)))) but is expected to have type - forall {α : Type.{u1}} [_inst_1 : Div.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.3764 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.3766 : Filter.{u1} α) => HDiv.hDiv.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHDiv.{u1} (Filter.{u1} α) (Filter.instDiv.{u1} α _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.3764 x._@.Mathlib.Order.Filter.Pointwise._hyg.3766) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.3779 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.3781 : Filter.{u1} α) => LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.instPartialOrderFilter.{u1} α))) x._@.Mathlib.Order.Filter.Pointwise._hyg.3779 x._@.Mathlib.Order.Filter.Pointwise._hyg.3781) + forall {α : Type.{u1}} [_inst_1 : Div.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.3763 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.3765 : Filter.{u1} α) => HDiv.hDiv.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHDiv.{u1} (Filter.{u1} α) (Filter.instDiv.{u1} α _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.3763 x._@.Mathlib.Order.Filter.Pointwise._hyg.3765) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.3778 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.3780 : Filter.{u1} α) => LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.instPartialOrderFilter.{u1} α))) x._@.Mathlib.Order.Filter.Pointwise._hyg.3778 x._@.Mathlib.Order.Filter.Pointwise._hyg.3780) Case conversion may be inaccurate. Consider using '#align filter.covariant_div Filter.covariant_divₓ'. -/ @[to_additive] instance covariant_div : CovariantClass (Filter α) (Filter α) (· / ·) (· ≤ ·) := @@ -775,7 +775,7 @@ instance covariant_div : CovariantClass (Filter α) (Filter α) (· / ·) (· lean 3 declaration is forall {α : Type.{u1}} [_inst_1 : Div.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Function.swap.{succ u1, succ u1, succ u1} (Filter.{u1} α) (Filter.{u1} α) (fun (ᾰ : Filter.{u1} α) (ᾰ : Filter.{u1} α) => Filter.{u1} α) (HDiv.hDiv.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHDiv.{u1} (Filter.{u1} α) (Filter.instDiv.{u1} α _inst_1)))) (LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.partialOrder.{u1} α)))) but is expected to have type - forall {α : Type.{u1}} [_inst_1 : Div.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Function.swap.{succ u1, succ u1, succ u1} (Filter.{u1} α) (Filter.{u1} α) (fun (ᾰ : Filter.{u1} α) (ᾰ : Filter.{u1} α) => Filter.{u1} α) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.3849 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.3851 : Filter.{u1} α) => HDiv.hDiv.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHDiv.{u1} (Filter.{u1} α) (Filter.instDiv.{u1} α _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.3849 x._@.Mathlib.Order.Filter.Pointwise._hyg.3851)) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.3864 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.3866 : Filter.{u1} α) => LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.instPartialOrderFilter.{u1} α))) x._@.Mathlib.Order.Filter.Pointwise._hyg.3864 x._@.Mathlib.Order.Filter.Pointwise._hyg.3866) + forall {α : Type.{u1}} [_inst_1 : Div.{u1} α], CovariantClass.{u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Function.swap.{succ u1, succ u1, succ u1} (Filter.{u1} α) (Filter.{u1} α) (fun (ᾰ : Filter.{u1} α) (ᾰ : Filter.{u1} α) => Filter.{u1} α) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.3848 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.3850 : Filter.{u1} α) => HDiv.hDiv.{u1, u1, u1} (Filter.{u1} α) (Filter.{u1} α) (Filter.{u1} α) (instHDiv.{u1} (Filter.{u1} α) (Filter.instDiv.{u1} α _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.3848 x._@.Mathlib.Order.Filter.Pointwise._hyg.3850)) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.3863 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.3865 : Filter.{u1} α) => LE.le.{u1} (Filter.{u1} α) (Preorder.toLE.{u1} (Filter.{u1} α) (PartialOrder.toPreorder.{u1} (Filter.{u1} α) (Filter.instPartialOrderFilter.{u1} α))) x._@.Mathlib.Order.Filter.Pointwise._hyg.3863 x._@.Mathlib.Order.Filter.Pointwise._hyg.3865) Case conversion may be inaccurate. Consider using '#align filter.covariant_swap_div Filter.covariant_swap_divₓ'. -/ @[to_additive] instance covariant_swap_div : CovariantClass (Filter α) (Filter α) (swap (· / ·)) (· ≤ ·) := @@ -1579,7 +1579,7 @@ theorem le_smul_iff : h ≤ f • g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : SMul.{u1, u2} α β], CovariantClass.{u1, u2} (Filter.{u1} α) (Filter.{u2} β) (SMul.smul.{u1, u2} (Filter.{u1} α) (Filter.{u2} β) (Filter.instSMul.{u1, u2} α β _inst_1)) (LE.le.{u2} (Filter.{u2} β) (Preorder.toLE.{u2} (Filter.{u2} β) (PartialOrder.toPreorder.{u2} (Filter.{u2} β) (Filter.partialOrder.{u2} β)))) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : SMul.{u1, u2} α β], CovariantClass.{u1, u2} (Filter.{u1} α) (Filter.{u2} β) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.8414 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.8416 : Filter.{u2} β) => HSMul.hSMul.{u1, u2, u2} (Filter.{u1} α) (Filter.{u2} β) (Filter.{u2} β) (instHSMul.{u1, u2} (Filter.{u1} α) (Filter.{u2} β) (Filter.instSMul.{u1, u2} α β _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.8414 x._@.Mathlib.Order.Filter.Pointwise._hyg.8416) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.8429 : Filter.{u2} β) (x._@.Mathlib.Order.Filter.Pointwise._hyg.8431 : Filter.{u2} β) => LE.le.{u2} (Filter.{u2} β) (Preorder.toLE.{u2} (Filter.{u2} β) (PartialOrder.toPreorder.{u2} (Filter.{u2} β) (Filter.instPartialOrderFilter.{u2} β))) x._@.Mathlib.Order.Filter.Pointwise._hyg.8429 x._@.Mathlib.Order.Filter.Pointwise._hyg.8431) + forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : SMul.{u1, u2} α β], CovariantClass.{u1, u2} (Filter.{u1} α) (Filter.{u2} β) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.8279 : Filter.{u1} α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.8281 : Filter.{u2} β) => HSMul.hSMul.{u1, u2, u2} (Filter.{u1} α) (Filter.{u2} β) (Filter.{u2} β) (instHSMul.{u1, u2} (Filter.{u1} α) (Filter.{u2} β) (Filter.instSMul.{u1, u2} α β _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.8279 x._@.Mathlib.Order.Filter.Pointwise._hyg.8281) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.8294 : Filter.{u2} β) (x._@.Mathlib.Order.Filter.Pointwise._hyg.8296 : Filter.{u2} β) => LE.le.{u2} (Filter.{u2} β) (Preorder.toLE.{u2} (Filter.{u2} β) (PartialOrder.toPreorder.{u2} (Filter.{u2} β) (Filter.instPartialOrderFilter.{u2} β))) x._@.Mathlib.Order.Filter.Pointwise._hyg.8294 x._@.Mathlib.Order.Filter.Pointwise._hyg.8296) Case conversion may be inaccurate. Consider using '#align filter.covariant_smul Filter.covariant_smulₓ'. -/ @[to_additive] instance covariant_smul : CovariantClass (Filter α) (Filter β) (· • ·) (· ≤ ·) := @@ -1613,7 +1613,7 @@ scoped[Pointwise] attribute [instance] Filter.instVSub lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : VSub.{u1, u2} α β] {f : Filter.{u2} β} {g : Filter.{u2} β}, Eq.{succ u1} (Filter.{u1} α) (Filter.map₂.{u2, u2, u1} β β α (VSub.vsub.{u1, u2} α β _inst_1) f g) (VSub.vsub.{u1, u2} (Filter.{u1} α) (Filter.{u2} β) (Filter.instVSub.{u1, u2} α β _inst_1) f g) but is expected to have type - forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : VSub.{u2, u1} α β] {f : Filter.{u1} β} {g : Filter.{u1} β}, Eq.{succ u2} (Filter.{u2} α) (Filter.map₂.{u1, u1, u2} β β α (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.8666 : β) (x._@.Mathlib.Order.Filter.Pointwise._hyg.8668 : β) => VSub.vsub.{u2, u1} α β _inst_1 x._@.Mathlib.Order.Filter.Pointwise._hyg.8666 x._@.Mathlib.Order.Filter.Pointwise._hyg.8668) f g) (VSub.vsub.{u2, u1} (Filter.{u2} α) (Filter.{u1} β) (Filter.instVSub.{u2, u1} α β _inst_1) f g) + forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : VSub.{u2, u1} α β] {f : Filter.{u1} β} {g : Filter.{u1} β}, Eq.{succ u2} (Filter.{u2} α) (Filter.map₂.{u1, u1, u2} β β α (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.8531 : β) (x._@.Mathlib.Order.Filter.Pointwise._hyg.8533 : β) => VSub.vsub.{u2, u1} α β _inst_1 x._@.Mathlib.Order.Filter.Pointwise._hyg.8531 x._@.Mathlib.Order.Filter.Pointwise._hyg.8533) f g) (VSub.vsub.{u2, u1} (Filter.{u2} α) (Filter.{u1} β) (Filter.instVSub.{u2, u1} α β _inst_1) f g) Case conversion may be inaccurate. Consider using '#align filter.map₂_vsub Filter.map₂_vsubₓ'. -/ @[simp] theorem map₂_vsub : map₂ (· -ᵥ ·) f g = f -ᵥ g := @@ -1710,7 +1710,7 @@ theorem NeBot.of_vsub_right : (f -ᵥ g : Filter α).ne_bot → g.ne_bot := lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : VSub.{u1, u2} α β] {g : Filter.{u2} β} {a : β}, Eq.{succ u1} (Filter.{u1} α) (VSub.vsub.{u1, u2} (Filter.{u1} α) (Filter.{u2} β) (Filter.instVSub.{u1, u2} α β _inst_1) (Pure.pure.{u2, u2} Filter.{u2} Filter.hasPure.{u2} β a) g) (Filter.map.{u2, u1} β α (VSub.vsub.{u1, u2} α β _inst_1 a) g) but is expected to have type - forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : VSub.{u2, u1} α β] {g : Filter.{u1} β} {a : β}, Eq.{succ u2} (Filter.{u2} α) (VSub.vsub.{u2, u1} (Filter.{u2} α) (Filter.{u1} β) (Filter.instVSub.{u2, u1} α β _inst_1) (Pure.pure.{u1, u1} Filter.{u1} Filter.instPureFilter.{u1} β a) g) (Filter.map.{u1, u2} β α ((fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.9232 : β) (x._@.Mathlib.Order.Filter.Pointwise._hyg.9234 : β) => VSub.vsub.{u2, u1} α β _inst_1 x._@.Mathlib.Order.Filter.Pointwise._hyg.9232 x._@.Mathlib.Order.Filter.Pointwise._hyg.9234) a) g) + forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : VSub.{u2, u1} α β] {g : Filter.{u1} β} {a : β}, Eq.{succ u2} (Filter.{u2} α) (VSub.vsub.{u2, u1} (Filter.{u2} α) (Filter.{u1} β) (Filter.instVSub.{u2, u1} α β _inst_1) (Pure.pure.{u1, u1} Filter.{u1} Filter.instPureFilter.{u1} β a) g) (Filter.map.{u1, u2} β α ((fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.9097 : β) (x._@.Mathlib.Order.Filter.Pointwise._hyg.9099 : β) => VSub.vsub.{u2, u1} α β _inst_1 x._@.Mathlib.Order.Filter.Pointwise._hyg.9097 x._@.Mathlib.Order.Filter.Pointwise._hyg.9099) a) g) Case conversion may be inaccurate. Consider using '#align filter.pure_vsub Filter.pure_vsubₓ'. -/ @[simp] theorem pure_vsub : (pure a : Filter β) -ᵥ g = g.map ((· -ᵥ ·) a) := @@ -1888,7 +1888,7 @@ theorem smul_filter_le_smul_filter (hf : f₁ ≤ f₂) : a • f₁ ≤ a • f lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : SMul.{u1, u2} α β], CovariantClass.{u1, u2} α (Filter.{u2} β) (SMul.smul.{u1, u2} α (Filter.{u2} β) (Filter.instSMulFilter.{u1, u2} α β _inst_1)) (LE.le.{u2} (Filter.{u2} β) (Preorder.toLE.{u2} (Filter.{u2} β) (PartialOrder.toPreorder.{u2} (Filter.{u2} β) (Filter.partialOrder.{u2} β)))) but is expected to have type - forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : SMul.{u1, u2} α β], CovariantClass.{u1, u2} α (Filter.{u2} β) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.10079 : α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.10081 : Filter.{u2} β) => HSMul.hSMul.{u1, u2, u2} α (Filter.{u2} β) (Filter.{u2} β) (instHSMul.{u1, u2} α (Filter.{u2} β) (Filter.instSMulFilter.{u1, u2} α β _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.10079 x._@.Mathlib.Order.Filter.Pointwise._hyg.10081) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.10094 : Filter.{u2} β) (x._@.Mathlib.Order.Filter.Pointwise._hyg.10096 : Filter.{u2} β) => LE.le.{u2} (Filter.{u2} β) (Preorder.toLE.{u2} (Filter.{u2} β) (PartialOrder.toPreorder.{u2} (Filter.{u2} β) (Filter.instPartialOrderFilter.{u2} β))) x._@.Mathlib.Order.Filter.Pointwise._hyg.10094 x._@.Mathlib.Order.Filter.Pointwise._hyg.10096) + forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : SMul.{u1, u2} α β], CovariantClass.{u1, u2} α (Filter.{u2} β) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.9944 : α) (x._@.Mathlib.Order.Filter.Pointwise._hyg.9946 : Filter.{u2} β) => HSMul.hSMul.{u1, u2, u2} α (Filter.{u2} β) (Filter.{u2} β) (instHSMul.{u1, u2} α (Filter.{u2} β) (Filter.instSMulFilter.{u1, u2} α β _inst_1)) x._@.Mathlib.Order.Filter.Pointwise._hyg.9944 x._@.Mathlib.Order.Filter.Pointwise._hyg.9946) (fun (x._@.Mathlib.Order.Filter.Pointwise._hyg.9959 : Filter.{u2} β) (x._@.Mathlib.Order.Filter.Pointwise._hyg.9961 : Filter.{u2} β) => LE.le.{u2} (Filter.{u2} β) (Preorder.toLE.{u2} (Filter.{u2} β) (PartialOrder.toPreorder.{u2} (Filter.{u2} β) (Filter.instPartialOrderFilter.{u2} β))) x._@.Mathlib.Order.Filter.Pointwise._hyg.9959 x._@.Mathlib.Order.Filter.Pointwise._hyg.9961) Case conversion may be inaccurate. Consider using '#align filter.covariant_smul_filter Filter.covariant_smul_filterₓ'. -/ @[to_additive] instance covariant_smul_filter : CovariantClass α (Filter β) (· • ·) (· ≤ ·) := diff --git a/Mathbin/Topology/Algebra/Group/Basic.lean b/Mathbin/Topology/Algebra/Group/Basic.lean index f884a297c0..02c51c1339 100644 --- a/Mathbin/Topology/Algebra/Group/Basic.lean +++ b/Mathbin/Topology/Algebra/Group/Basic.lean @@ -118,11 +118,11 @@ lean 3 declaration is but is expected to have type forall {G : Type.{u1}} [_inst_1 : TopologicalSpace.{u1} G] [_inst_2 : Group.{u1} G] [_inst_3 : ContinuousMul.{u1} G _inst_1 (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_2))))] {U : Set.{u1} G}, (IsOpen.{u1} G _inst_1 U) -> (forall (x : G), IsOpen.{u1} G _inst_1 (leftCoset.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_2)))) x U)) Case conversion may be inaccurate. Consider using '#align is_open.left_coset IsOpen.leftCosetₓ'. -/ -@[to_additive IsOpen.left_add_coset] +@[to_additive IsOpen.left_addCoset] theorem IsOpen.leftCoset {U : Set G} (h : IsOpen U) (x : G) : IsOpen (leftCoset x U) := isOpenMap_mul_left x _ h #align is_open.left_coset IsOpen.leftCoset -#align is_open.left_add_coset IsOpen.left_add_coset +#align is_open.left_add_coset IsOpen.left_addCoset /- warning: is_closed_map_mul_left -> isClosedMap_mul_left is a dubious translation: lean 3 declaration is @@ -142,11 +142,11 @@ lean 3 declaration is but is expected to have type forall {G : Type.{u1}} [_inst_1 : TopologicalSpace.{u1} G] [_inst_2 : Group.{u1} G] [_inst_3 : ContinuousMul.{u1} G _inst_1 (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_2))))] {U : Set.{u1} G}, (IsClosed.{u1} G _inst_1 U) -> (forall (x : G), IsClosed.{u1} G _inst_1 (leftCoset.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_2)))) x U)) Case conversion may be inaccurate. Consider using '#align is_closed.left_coset IsClosed.leftCosetₓ'. -/ -@[to_additive IsClosed.left_add_coset] +@[to_additive IsClosed.left_addCoset] theorem IsClosed.leftCoset {U : Set G} (h : IsClosed U) (x : G) : IsClosed (leftCoset x U) := isClosedMap_mul_left x _ h #align is_closed.left_coset IsClosed.leftCoset -#align is_closed.left_add_coset IsClosed.left_add_coset +#align is_closed.left_add_coset IsClosed.left_addCoset /- warning: homeomorph.mul_right -> Homeomorph.mulRight is a dubious translation: lean 3 declaration is @@ -207,11 +207,11 @@ lean 3 declaration is but is expected to have type forall {G : Type.{u1}} [_inst_1 : TopologicalSpace.{u1} G] [_inst_2 : Group.{u1} G] [_inst_3 : ContinuousMul.{u1} G _inst_1 (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_2))))] {U : Set.{u1} G}, (IsOpen.{u1} G _inst_1 U) -> (forall (x : G), IsOpen.{u1} G _inst_1 (rightCoset.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_2)))) U x)) Case conversion may be inaccurate. Consider using '#align is_open.right_coset IsOpen.rightCosetₓ'. -/ -@[to_additive IsOpen.right_add_coset] +@[to_additive IsOpen.right_addCoset] theorem IsOpen.rightCoset {U : Set G} (h : IsOpen U) (x : G) : IsOpen (rightCoset U x) := isOpenMap_mul_right x _ h #align is_open.right_coset IsOpen.rightCoset -#align is_open.right_add_coset IsOpen.right_add_coset +#align is_open.right_add_coset IsOpen.right_addCoset /- warning: is_closed_map_mul_right -> isClosedMap_mul_right is a dubious translation: lean 3 declaration is @@ -231,11 +231,11 @@ lean 3 declaration is but is expected to have type forall {G : Type.{u1}} [_inst_1 : TopologicalSpace.{u1} G] [_inst_2 : Group.{u1} G] [_inst_3 : ContinuousMul.{u1} G _inst_1 (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_2))))] {U : Set.{u1} G}, (IsClosed.{u1} G _inst_1 U) -> (forall (x : G), IsClosed.{u1} G _inst_1 (rightCoset.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_2)))) U x)) Case conversion may be inaccurate. Consider using '#align is_closed.right_coset IsClosed.rightCosetₓ'. -/ -@[to_additive IsClosed.right_add_coset] +@[to_additive IsClosed.right_addCoset] theorem IsClosed.rightCoset {U : Set G} (h : IsClosed U) (x : G) : IsClosed (rightCoset U x) := isClosedMap_mul_right x _ h #align is_closed.right_coset IsClosed.rightCoset -#align is_closed.right_add_coset IsClosed.right_add_coset +#align is_closed.right_add_coset IsClosed.right_addCoset /- warning: discrete_topology_of_open_singleton_one -> discreteTopology_of_open_singleton_one is a dubious translation: lean 3 declaration is @@ -721,16 +721,16 @@ instance AddGroup.continuousConstSMul_int {A} [AddGroup A] [TopologicalSpace A] #align add_group.has_continuous_const_smul_int AddGroup.continuousConstSMul_int -/ -/- warning: add_group.has_continuous_smul_int -> AddGroup.continuousSmul_int is a dubious translation: +/- warning: add_group.has_continuous_smul_int -> AddGroup.continuousSMul_int is a dubious translation: lean 3 declaration is forall {A : Type.{u1}} [_inst_5 : AddGroup.{u1} A] [_inst_6 : TopologicalSpace.{u1} A] [_inst_7 : TopologicalAddGroup.{u1} A _inst_6 _inst_5], ContinuousSMul.{0, u1} Int A (SubNegMonoid.SMulInt.{u1} A (AddGroup.toSubNegMonoid.{u1} A _inst_5)) Int.topologicalSpace _inst_6 but is expected to have type forall {A : Type.{u1}} [_inst_5 : AddGroup.{u1} A] [_inst_6 : TopologicalSpace.{u1} A] [_inst_7 : TopologicalAddGroup.{u1} A _inst_6 _inst_5], ContinuousSMul.{0, u1} Int A (SubNegMonoid.SMulInt.{u1} A (AddGroup.toSubNegMonoid.{u1} A _inst_5)) instTopologicalSpaceInt _inst_6 -Case conversion may be inaccurate. Consider using '#align add_group.has_continuous_smul_int AddGroup.continuousSmul_intₓ'. -/ -instance AddGroup.continuousSmul_int {A} [AddGroup A] [TopologicalSpace A] [TopologicalAddGroup A] : +Case conversion may be inaccurate. Consider using '#align add_group.has_continuous_smul_int AddGroup.continuousSMul_intₓ'. -/ +instance AddGroup.continuousSMul_int {A} [AddGroup A] [TopologicalSpace A] [TopologicalAddGroup A] : ContinuousSMul ℤ A := ⟨continuous_uncurry_of_discreteTopology continuous_zsmul⟩ -#align add_group.has_continuous_smul_int AddGroup.continuousSmul_int +#align add_group.has_continuous_smul_int AddGroup.continuousSMul_int #print Continuous.zpow /- @[continuity, to_additive] @@ -1381,7 +1381,7 @@ theorem ContinuousInv.of_nhds_one {G : Type _} [Group G] [TopologicalSpace G] lean 3 declaration is forall {G : Type.{u1}} [_inst_5 : Group.{u1} G] [_inst_6 : TopologicalSpace.{u1} G], (Filter.Tendsto.{u1, u1} (Prod.{u1, u1} G G) G (Function.uncurry.{u1, u1, u1} G G G (HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))))) (Filter.prod.{u1, u1} G G (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))))))) -> (Filter.Tendsto.{u1, u1} G G (fun (x : G) => Inv.inv.{u1} G (DivInvMonoid.toHasInv.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)) x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x₀ x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)))))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x x₀) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)))))))))) -> (TopologicalGroup.{u1} G _inst_6 _inst_5) but is expected to have type - forall {G : Type.{u1}} [_inst_5 : Group.{u1} G] [_inst_6 : TopologicalSpace.{u1} G], (Filter.Tendsto.{u1, u1} (Prod.{u1, u1} G G) G (Function.uncurry.{u1, u1, u1} G G G (fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7129 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7131 : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7129 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7131)) (Filter.prod.{u1, u1} G G (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (Filter.Tendsto.{u1, u1} G G (fun (x : G) => Inv.inv.{u1} G (InvOneClass.toInv.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))) x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x₀ x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x x₀) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))))) -> (TopologicalGroup.{u1} G _inst_6 _inst_5) + forall {G : Type.{u1}} [_inst_5 : Group.{u1} G] [_inst_6 : TopologicalSpace.{u1} G], (Filter.Tendsto.{u1, u1} (Prod.{u1, u1} G G) G (Function.uncurry.{u1, u1, u1} G G G (fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7130 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7132 : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7130 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7132)) (Filter.prod.{u1, u1} G G (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (Filter.Tendsto.{u1, u1} G G (fun (x : G) => Inv.inv.{u1} G (InvOneClass.toInv.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))) x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x₀ x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x x₀) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))))) -> (TopologicalGroup.{u1} G _inst_6 _inst_5) Case conversion may be inaccurate. Consider using '#align topological_group.of_nhds_one' TopologicalGroup.of_nhds_one'ₓ'. -/ @[to_additive] theorem TopologicalGroup.of_nhds_one' {G : Type u} [Group G] [TopologicalSpace G] @@ -1404,7 +1404,7 @@ theorem TopologicalGroup.of_nhds_one' {G : Type u} [Group G] [TopologicalSpace G lean 3 declaration is forall {G : Type.{u1}} [_inst_5 : Group.{u1} G] [_inst_6 : TopologicalSpace.{u1} G], (Filter.Tendsto.{u1, u1} (Prod.{u1, u1} G G) G (Function.uncurry.{u1, u1, u1} G G G (HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))))) (Filter.prod.{u1, u1} G G (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))))))) -> (Filter.Tendsto.{u1, u1} G G (fun (x : G) => Inv.inv.{u1} G (DivInvMonoid.toHasInv.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)) x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x₀ x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)))))))))) -> (forall (x₀ : G), Filter.Tendsto.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) (HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x₀ x) (Inv.inv.{u1} G (DivInvMonoid.toHasInv.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)) x₀)) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))))))) -> (TopologicalGroup.{u1} G _inst_6 _inst_5) but is expected to have type - forall {G : Type.{u1}} [_inst_5 : Group.{u1} G] [_inst_6 : TopologicalSpace.{u1} G], (Filter.Tendsto.{u1, u1} (Prod.{u1, u1} G G) G (Function.uncurry.{u1, u1, u1} G G G (fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7417 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7419 : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7417 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7419)) (Filter.prod.{u1, u1} G G (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (Filter.Tendsto.{u1, u1} G G (fun (x : G) => Inv.inv.{u1} G (InvOneClass.toInv.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))) x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x₀ x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))))) -> (forall (x₀ : G), Filter.Tendsto.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) (HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x₀ x) (Inv.inv.{u1} G (InvOneClass.toInv.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))) x₀)) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (TopologicalGroup.{u1} G _inst_6 _inst_5) + forall {G : Type.{u1}} [_inst_5 : Group.{u1} G] [_inst_6 : TopologicalSpace.{u1} G], (Filter.Tendsto.{u1, u1} (Prod.{u1, u1} G G) G (Function.uncurry.{u1, u1, u1} G G G (fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7418 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7420 : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7418 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7420)) (Filter.prod.{u1, u1} G G (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (Filter.Tendsto.{u1, u1} G G (fun (x : G) => Inv.inv.{u1} G (InvOneClass.toInv.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))) x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x₀ x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))))) -> (forall (x₀ : G), Filter.Tendsto.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) (HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G _inst_5))))) x₀ x) (Inv.inv.{u1} G (InvOneClass.toInv.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))) x₀)) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (Group.toDivisionMonoid.{u1} G _inst_5)))))))) -> (TopologicalGroup.{u1} G _inst_6 _inst_5) Case conversion may be inaccurate. Consider using '#align topological_group.of_nhds_one TopologicalGroup.of_nhds_oneₓ'. -/ @[to_additive] theorem TopologicalGroup.of_nhds_one {G : Type u} [Group G] [TopologicalSpace G] @@ -1430,7 +1430,7 @@ theorem TopologicalGroup.of_nhds_one {G : Type u} [Group G] [TopologicalSpace G] lean 3 declaration is forall {G : Type.{u1}} [_inst_5 : CommGroup.{u1} G] [_inst_6 : TopologicalSpace.{u1} G], (Filter.Tendsto.{u1, u1} (Prod.{u1, u1} G G) G (Function.uncurry.{u1, u1, u1} G G G (HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5)))))))) (Filter.prod.{u1, u1} G G (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5))))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5)))))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5)))))))))) -> (Filter.Tendsto.{u1, u1} G G (fun (x : G) => Inv.inv.{u1} G (DivInvMonoid.toHasInv.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5))) x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5))))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5)))))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toHasMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5)))))) x₀ x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (OfNat.mk.{u1} G 1 (One.one.{u1} G (MulOneClass.toHasOne.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5))))))))))) -> (TopologicalGroup.{u1} G _inst_6 (CommGroup.toGroup.{u1} G _inst_5)) but is expected to have type - forall {G : Type.{u1}} [_inst_5 : CommGroup.{u1} G] [_inst_6 : TopologicalSpace.{u1} G], (Filter.Tendsto.{u1, u1} (Prod.{u1, u1} G G) G (Function.uncurry.{u1, u1, u1} G G G (fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7733 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7735 : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5)))))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7733 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7735)) (Filter.prod.{u1, u1} G G (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5))))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5))))))))) -> (Filter.Tendsto.{u1, u1} G G (fun (x : G) => Inv.inv.{u1} G (InvOneClass.toInv.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5))))) x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5))))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5)))))) x₀ x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5)))))))))) -> (TopologicalGroup.{u1} G _inst_6 (CommGroup.toGroup.{u1} G _inst_5)) + forall {G : Type.{u1}} [_inst_5 : CommGroup.{u1} G] [_inst_6 : TopologicalSpace.{u1} G], (Filter.Tendsto.{u1, u1} (Prod.{u1, u1} G G) G (Function.uncurry.{u1, u1, u1} G G G (fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7736 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7738 : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5)))))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7736 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.7738)) (Filter.prod.{u1, u1} G G (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5))))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5))))))))) -> (Filter.Tendsto.{u1, u1} G G (fun (x : G) => Inv.inv.{u1} G (InvOneClass.toInv.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5))))) x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5)))))))) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5))))))))) -> (forall (x₀ : G), Eq.{succ u1} (Filter.{u1} G) (nhds.{u1} G _inst_6 x₀) (Filter.map.{u1, u1} G G (fun (x : G) => HMul.hMul.{u1, u1, u1} G G G (instHMul.{u1} G (MulOneClass.toMul.{u1} G (Monoid.toMulOneClass.{u1} G (DivInvMonoid.toMonoid.{u1} G (Group.toDivInvMonoid.{u1} G (CommGroup.toGroup.{u1} G _inst_5)))))) x₀ x) (nhds.{u1} G _inst_6 (OfNat.ofNat.{u1} G 1 (One.toOfNat1.{u1} G (InvOneClass.toOne.{u1} G (DivInvOneMonoid.toInvOneClass.{u1} G (DivisionMonoid.toDivInvOneMonoid.{u1} G (DivisionCommMonoid.toDivisionMonoid.{u1} G (CommGroup.toDivisionCommMonoid.{u1} G _inst_5)))))))))) -> (TopologicalGroup.{u1} G _inst_6 (CommGroup.toGroup.{u1} G _inst_5)) Case conversion may be inaccurate. Consider using '#align topological_group.of_comm_of_nhds_one TopologicalGroup.of_comm_of_nhds_oneₓ'. -/ @[to_additive] theorem TopologicalGroup.of_comm_of_nhds_one {G : Type u} [CommGroup G] [TopologicalSpace G] @@ -1697,7 +1697,7 @@ def Homeomorph.divLeft (x : G) : G ≃ₜ G := lean 3 declaration is forall {G : Type.{u1}} [_inst_1 : Group.{u1} G] [_inst_2 : TopologicalSpace.{u1} G] [_inst_3 : TopologicalGroup.{u1} G _inst_2 _inst_1] (a : G), IsOpenMap.{u1, u1} G G _inst_2 _inst_2 (HDiv.hDiv.{u1, u1, u1} G G G (instHDiv.{u1} G (DivInvMonoid.toHasDiv.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1))) a) but is expected to have type - forall {G : Type.{u1}} [_inst_1 : Group.{u1} G] [_inst_2 : TopologicalSpace.{u1} G] [_inst_3 : TopologicalGroup.{u1} G _inst_2 _inst_1] (a : G), IsOpenMap.{u1, u1} G G _inst_2 _inst_2 ((fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9546 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9548 : G) => HDiv.hDiv.{u1, u1, u1} G G G (instHDiv.{u1} G (DivInvMonoid.toDiv.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9546 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9548) a) + forall {G : Type.{u1}} [_inst_1 : Group.{u1} G] [_inst_2 : TopologicalSpace.{u1} G] [_inst_3 : TopologicalGroup.{u1} G _inst_2 _inst_1] (a : G), IsOpenMap.{u1, u1} G G _inst_2 _inst_2 ((fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9549 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9551 : G) => HDiv.hDiv.{u1, u1, u1} G G G (instHDiv.{u1} G (DivInvMonoid.toDiv.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9549 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9551) a) Case conversion may be inaccurate. Consider using '#align is_open_map_div_left isOpenMap_div_leftₓ'. -/ @[to_additive] theorem isOpenMap_div_left (a : G) : IsOpenMap ((· / ·) a) := @@ -1709,7 +1709,7 @@ theorem isOpenMap_div_left (a : G) : IsOpenMap ((· / ·) a) := lean 3 declaration is forall {G : Type.{u1}} [_inst_1 : Group.{u1} G] [_inst_2 : TopologicalSpace.{u1} G] [_inst_3 : TopologicalGroup.{u1} G _inst_2 _inst_1] (a : G), IsClosedMap.{u1, u1} G G _inst_2 _inst_2 (HDiv.hDiv.{u1, u1, u1} G G G (instHDiv.{u1} G (DivInvMonoid.toHasDiv.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1))) a) but is expected to have type - forall {G : Type.{u1}} [_inst_1 : Group.{u1} G] [_inst_2 : TopologicalSpace.{u1} G] [_inst_3 : TopologicalGroup.{u1} G _inst_2 _inst_1] (a : G), IsClosedMap.{u1, u1} G G _inst_2 _inst_2 ((fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9588 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9590 : G) => HDiv.hDiv.{u1, u1, u1} G G G (instHDiv.{u1} G (DivInvMonoid.toDiv.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9588 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9590) a) + forall {G : Type.{u1}} [_inst_1 : Group.{u1} G] [_inst_2 : TopologicalSpace.{u1} G] [_inst_3 : TopologicalGroup.{u1} G _inst_2 _inst_1] (a : G), IsClosedMap.{u1, u1} G G _inst_2 _inst_2 ((fun (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9591 : G) (x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9593 : G) => HDiv.hDiv.{u1, u1, u1} G G G (instHDiv.{u1} G (DivInvMonoid.toDiv.{u1} G (Group.toDivInvMonoid.{u1} G _inst_1))) x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9591 x._@.Mathlib.Topology.Algebra.Group.Basic._hyg.9593) a) Case conversion may be inaccurate. Consider using '#align is_closed_map_div_left isClosedMap_div_leftₓ'. -/ @[to_additive] theorem isClosedMap_div_left (a : G) : IsClosedMap ((· / ·) a) := diff --git a/Mathbin/Topology/List.lean b/Mathbin/Topology/List.lean index 67c378daae..19cf35143c 100644 --- a/Mathbin/Topology/List.lean +++ b/Mathbin/Topology/List.lean @@ -194,16 +194,16 @@ theorem continuousAt_length : ∀ l : List α, ContinuousAt List.length l := refine' tendsto.comp ih tendsto_snd #align list.continuous_at_length List.continuousAt_length -/- warning: list.tendsto_insert_nth' -> List.tendsto_insert_nth' is a dubious translation: +/- warning: list.tendsto_insert_nth' -> List.tendsto_insertNth' is a dubious translation: lean 3 declaration is forall {α : Type.{u1}} [_inst_1 : TopologicalSpace.{u1} α] {a : α} {n : Nat} {l : List.{u1} α}, Filter.Tendsto.{u1, u1} (Prod.{u1, u1} α (List.{u1} α)) (List.{u1} α) (fun (p : Prod.{u1, u1} α (List.{u1} α)) => List.insertNth.{u1} α n (Prod.fst.{u1, u1} α (List.{u1} α) p) (Prod.snd.{u1, u1} α (List.{u1} α) p)) (Filter.prod.{u1, u1} α (List.{u1} α) (nhds.{u1} α _inst_1 a) (nhds.{u1} (List.{u1} α) (List.topologicalSpace.{u1} α _inst_1) l)) (nhds.{u1} (List.{u1} α) (List.topologicalSpace.{u1} α _inst_1) (List.insertNth.{u1} α n a l)) but is expected to have type forall {α : Type.{u1}} [_inst_1 : TopologicalSpace.{u1} α] {a : α} {n : Nat} {l : List.{u1} α}, Filter.Tendsto.{u1, u1} (Prod.{u1, u1} α (List.{u1} α)) (List.{u1} α) (fun (p : Prod.{u1, u1} α (List.{u1} α)) => List.insertNth.{u1} α n (Prod.fst.{u1, u1} α (List.{u1} α) p) (Prod.snd.{u1, u1} α (List.{u1} α) p)) (Filter.prod.{u1, u1} α (List.{u1} α) (nhds.{u1} α _inst_1 a) (nhds.{u1} (List.{u1} α) (instTopologicalSpaceList.{u1} α _inst_1) l)) (nhds.{u1} (List.{u1} α) (instTopologicalSpaceList.{u1} α _inst_1) (List.insertNth.{u1} α n a l)) -Case conversion may be inaccurate. Consider using '#align list.tendsto_insert_nth' List.tendsto_insert_nth'ₓ'. -/ +Case conversion may be inaccurate. Consider using '#align list.tendsto_insert_nth' List.tendsto_insertNth'ₓ'. -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -theorem tendsto_insert_nth' {a : α} : +theorem tendsto_insertNth' {a : α} : ∀ {n : ℕ} {l : List α}, Tendsto (fun p : α × List α => insertNth n p.1 p.2) (𝓝 a ×ᶠ 𝓝 l) (𝓝 (insertNth n a l)) | 0, l => tendsto_cons @@ -219,7 +219,7 @@ theorem tendsto_insert_nth' {a : α} : exact (tendsto_fst.comp tendsto_snd).cons ((@tendsto_insert_nth' n l).comp <| tendsto_fst.prod_mk <| tendsto_snd.comp tendsto_snd) -#align list.tendsto_insert_nth' List.tendsto_insert_nth' +#align list.tendsto_insert_nth' List.tendsto_insertNth' /- warning: list.tendsto_insert_nth -> List.tendsto_insertNth is a dubious translation: lean 3 declaration is @@ -230,7 +230,7 @@ Case conversion may be inaccurate. Consider using '#align list.tendsto_insert_nt theorem tendsto_insertNth {β} {n : ℕ} {a : α} {l : List α} {f : β → α} {g : β → List α} {b : Filter β} (hf : Tendsto f b (𝓝 a)) (hg : Tendsto g b (𝓝 l)) : Tendsto (fun b : β => insertNth n (f b) (g b)) b (𝓝 (insertNth n a l)) := - tendsto_insert_nth'.comp (Tendsto.prod_mk hf hg) + tendsto_insertNth'.comp (Tendsto.prod_mk hf hg) #align list.tendsto_insert_nth List.tendsto_insertNth /- warning: list.continuous_insert_nth -> List.continuous_insertNth is a dubious translation: @@ -339,17 +339,17 @@ theorem tendsto_insertNth {n : ℕ} {i : Fin (n + 1)} {a : α} : exact List.tendsto_insertNth tendsto_fst (tendsto.comp continuousAt_subtype_val tendsto_snd : _) #align vector.tendsto_insert_nth Vector.tendsto_insertNth -/- warning: vector.continuous_insert_nth' -> Vector.continuous_insert_nth' is a dubious translation: +/- warning: vector.continuous_insert_nth' -> Vector.continuous_insertNth' is a dubious translation: lean 3 declaration is forall {α : Type.{u1}} [_inst_1 : TopologicalSpace.{u1} α] {n : Nat} {i : Fin (HAdd.hAdd.{0, 0, 0} Nat Nat Nat (instHAdd.{0} Nat Nat.hasAdd) n (OfNat.ofNat.{0} Nat 1 (OfNat.mk.{0} Nat 1 (One.one.{0} Nat Nat.hasOne))))}, Continuous.{u1, u1} (Prod.{u1, u1} α (Vector.{u1} α n)) (Vector.{u1} α (HAdd.hAdd.{0, 0, 0} Nat Nat Nat (instHAdd.{0} Nat Nat.hasAdd) n (OfNat.ofNat.{0} Nat 1 (OfNat.mk.{0} Nat 1 (One.one.{0} Nat Nat.hasOne))))) (Prod.topologicalSpace.{u1, u1} α (Vector.{u1} α n) _inst_1 (Vector.topologicalSpace.{u1} α _inst_1 n)) (Vector.topologicalSpace.{u1} α _inst_1 (HAdd.hAdd.{0, 0, 0} Nat Nat Nat (instHAdd.{0} Nat Nat.hasAdd) n (OfNat.ofNat.{0} Nat 1 (OfNat.mk.{0} Nat 1 (One.one.{0} Nat Nat.hasOne))))) (fun (p : Prod.{u1, u1} α (Vector.{u1} α n)) => Vector.insertNth.{u1} n α (Prod.fst.{u1, u1} α (Vector.{u1} α n) p) i (Prod.snd.{u1, u1} α (Vector.{u1} α n) p)) but is expected to have type forall {α : Type.{u1}} [_inst_1 : TopologicalSpace.{u1} α] {n : Nat} {i : Fin (HAdd.hAdd.{0, 0, 0} Nat Nat Nat (instHAdd.{0} Nat instAddNat) n (OfNat.ofNat.{0} Nat 1 (instOfNatNat 1)))}, Continuous.{u1, u1} (Prod.{u1, u1} α (Vector.{u1} α n)) (Vector.{u1} α (HAdd.hAdd.{0, 0, 0} Nat Nat Nat (instHAdd.{0} Nat instAddNat) n (OfNat.ofNat.{0} Nat 1 (instOfNatNat 1)))) (instTopologicalSpaceProd.{u1, u1} α (Vector.{u1} α n) _inst_1 (Vector.instTopologicalSpaceVector.{u1} α _inst_1 n)) (Vector.instTopologicalSpaceVector.{u1} α _inst_1 (HAdd.hAdd.{0, 0, 0} Nat Nat Nat (instHAdd.{0} Nat instAddNat) n (OfNat.ofNat.{0} Nat 1 (instOfNatNat 1)))) (fun (p : Prod.{u1, u1} α (Vector.{u1} α n)) => Vector.insertNth.{u1} n α (Prod.fst.{u1, u1} α (Vector.{u1} α n) p) i (Prod.snd.{u1, u1} α (Vector.{u1} α n) p)) -Case conversion may be inaccurate. Consider using '#align vector.continuous_insert_nth' Vector.continuous_insert_nth'ₓ'. -/ -theorem continuous_insert_nth' {n : ℕ} {i : Fin (n + 1)} : +Case conversion may be inaccurate. Consider using '#align vector.continuous_insert_nth' Vector.continuous_insertNth'ₓ'. -/ +theorem continuous_insertNth' {n : ℕ} {i : Fin (n + 1)} : Continuous fun p : α × Vector α n => insertNth p.1 i p.2 := continuous_iff_continuousAt.mpr fun ⟨a, l⟩ => by rw [ContinuousAt, nhds_prod_eq] <;> exact tendsto_insert_nth -#align vector.continuous_insert_nth' Vector.continuous_insert_nth' +#align vector.continuous_insert_nth' Vector.continuous_insertNth' /- warning: vector.continuous_insert_nth -> Vector.continuous_insertNth is a dubious translation: lean 3 declaration is @@ -359,7 +359,7 @@ but is expected to have type Case conversion may be inaccurate. Consider using '#align vector.continuous_insert_nth Vector.continuous_insertNthₓ'. -/ theorem continuous_insertNth {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → Vector α n} (hf : Continuous f) (hg : Continuous g) : Continuous fun b => insertNth (f b) i (g b) := - continuous_insert_nth'.comp (hf.prod_mk hg : _) + continuous_insertNth'.comp (hf.prod_mk hg : _) #align vector.continuous_insert_nth Vector.continuous_insertNth /- warning: vector.continuous_at_remove_nth -> Vector.continuousAt_removeNth is a dubious translation: diff --git a/Mathbin/Topology/Separation.lean b/Mathbin/Topology/Separation.lean index 21a246721c..e0553d9221 100644 --- a/Mathbin/Topology/Separation.lean +++ b/Mathbin/Topology/Separation.lean @@ -1940,7 +1940,7 @@ theorem Bornology.relativelyCompact_eq_inCompact [T2Space α] : lean 3 declaration is forall {α : Type.{u1}} [_inst_1 : TopologicalSpace.{u1} α] [_inst_4 : T2Space.{u1} α _inst_1] {ι : Type.{u2}} [_inst_5 : Nonempty.{succ u2} ι] {V : ι -> (Set.{u1} α)}, (Directed.{u1, succ u2} (Set.{u1} α) ι (Superset.{u1} (Set.{u1} α) (Set.hasSubset.{u1} α)) V) -> (forall (i : ι), IsCompact.{u1} α _inst_1 (V i)) -> (forall {U : Set.{u1} α}, (forall (x : α), (Membership.Mem.{u1, u1} α (Set.{u1} α) (Set.hasMem.{u1} α) x (Set.interᵢ.{u1, succ u2} α ι (fun (i : ι) => V i))) -> (Membership.Mem.{u1, u1} (Set.{u1} α) (Filter.{u1} α) (Filter.hasMem.{u1} α) U (nhds.{u1} α _inst_1 x))) -> (Exists.{succ u2} ι (fun (i : ι) => HasSubset.Subset.{u1} (Set.{u1} α) (Set.hasSubset.{u1} α) (V i) U))) but is expected to have type - forall {α : Type.{u2}} [_inst_1 : TopologicalSpace.{u2} α] [_inst_4 : T2Space.{u2} α _inst_1] {ι : Type.{u1}} [_inst_5 : Nonempty.{succ u1} ι] {V : ι -> (Set.{u2} α)}, (Directed.{u2, succ u1} (Set.{u2} α) ι (fun (x._@.Mathlib.Topology.Separation._hyg.13084 : Set.{u2} α) (x._@.Mathlib.Topology.Separation._hyg.13086 : Set.{u2} α) => Superset.{u2} (Set.{u2} α) (Set.instHasSubsetSet.{u2} α) x._@.Mathlib.Topology.Separation._hyg.13084 x._@.Mathlib.Topology.Separation._hyg.13086) V) -> (forall (i : ι), IsCompact.{u2} α _inst_1 (V i)) -> (forall {U : Set.{u2} α}, (forall (x : α), (Membership.mem.{u2, u2} α (Set.{u2} α) (Set.instMembershipSet.{u2} α) x (Set.interᵢ.{u2, succ u1} α ι (fun (i : ι) => V i))) -> (Membership.mem.{u2, u2} (Set.{u2} α) (Filter.{u2} α) (instMembershipSetFilter.{u2} α) U (nhds.{u2} α _inst_1 x))) -> (Exists.{succ u1} ι (fun (i : ι) => HasSubset.Subset.{u2} (Set.{u2} α) (Set.instHasSubsetSet.{u2} α) (V i) U))) + forall {α : Type.{u2}} [_inst_1 : TopologicalSpace.{u2} α] [_inst_4 : T2Space.{u2} α _inst_1] {ι : Type.{u1}} [_inst_5 : Nonempty.{succ u1} ι] {V : ι -> (Set.{u2} α)}, (Directed.{u2, succ u1} (Set.{u2} α) ι (fun (x._@.Mathlib.Topology.Separation._hyg.13060 : Set.{u2} α) (x._@.Mathlib.Topology.Separation._hyg.13062 : Set.{u2} α) => Superset.{u2} (Set.{u2} α) (Set.instHasSubsetSet.{u2} α) x._@.Mathlib.Topology.Separation._hyg.13060 x._@.Mathlib.Topology.Separation._hyg.13062) V) -> (forall (i : ι), IsCompact.{u2} α _inst_1 (V i)) -> (forall {U : Set.{u2} α}, (forall (x : α), (Membership.mem.{u2, u2} α (Set.{u2} α) (Set.instMembershipSet.{u2} α) x (Set.interᵢ.{u2, succ u1} α ι (fun (i : ι) => V i))) -> (Membership.mem.{u2, u2} (Set.{u2} α) (Filter.{u2} α) (instMembershipSetFilter.{u2} α) U (nhds.{u2} α _inst_1 x))) -> (Exists.{succ u1} ι (fun (i : ι) => HasSubset.Subset.{u2} (Set.{u2} α) (Set.instHasSubsetSet.{u2} α) (V i) U))) Case conversion may be inaccurate. Consider using '#align exists_subset_nhds_of_is_compact exists_subset_nhds_of_isCompactₓ'. -/ /-- If `V : ι → set α` is a decreasing family of compact sets then any neighborhood of `⋂ i, V i` contains some `V i`. This is a version of `exists_subset_nhd_of_compact'` where we diff --git a/lake-manifest.json b/lake-manifest.json index 6bcbc02a32..357abfe478 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -4,15 +4,15 @@ [{"git": {"url": "https://github.com/leanprover-community/lean3port.git", "subDir?": null, - "rev": "842162938b13898101b6567a3d365d3dd3f35154", + "rev": "3dd3e2c5bcbd00b7f7597f118389c9d4ddbdeda7", "name": "lean3port", - "inputRev?": "842162938b13898101b6567a3d365d3dd3f35154"}}, + "inputRev?": "3dd3e2c5bcbd00b7f7597f118389c9d4ddbdeda7"}}, {"git": {"url": "https://github.com/leanprover-community/mathlib4.git", "subDir?": null, - "rev": "b000aad8f4c9c180743c7d85d46ea376b48263bb", + "rev": "7392084b8235417c0a62e7c1f0753e5bae4b7851", "name": "mathlib", - "inputRev?": "b000aad8f4c9c180743c7d85d46ea376b48263bb"}}, + "inputRev?": "7392084b8235417c0a62e7c1f0753e5bae4b7851"}}, {"git": {"url": "https://github.com/gebner/quote4", "subDir?": null, diff --git a/lakefile.lean b/lakefile.lean index c837696cb5..59a68be04c 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -4,7 +4,7 @@ open Lake DSL System -- Usually the `tag` will be of the form `nightly-2021-11-22`. -- If you would like to use an artifact from a PR build, -- it will be of the form `pr-branchname-sha`. -def tag : String := "nightly-2023-04-17-08" +def tag : String := "nightly-2023-04-17-10" def releaseRepo : String := "leanprover-community/mathport" def oleanTarName : String := "mathlib3-binport.tar.gz" @@ -38,7 +38,7 @@ target fetchOleans (_pkg : Package) : Unit := do untarReleaseArtifact releaseRepo tag oleanTarName libDir return .nil -require lean3port from git "https://github.com/leanprover-community/lean3port.git"@"842162938b13898101b6567a3d365d3dd3f35154" +require lean3port from git "https://github.com/leanprover-community/lean3port.git"@"3dd3e2c5bcbd00b7f7597f118389c9d4ddbdeda7" @[default_target] lean_lib Mathbin where