From af7dd13c1166428d6c4dcdd496cfd51923dd152b Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Thu, 30 Mar 2023 02:56:05 +0000 Subject: [PATCH] bump to nightly-2023-03-30-02 mathlib commit https://github.com/leanprover-community/mathlib/commit/b685f506164f8d17a6404048bc4d696739c5d976 --- Mathbin/Algebra/ContinuedFractions/Basic.lean | 17 +- .../Computation/Approximations.lean | 6 +- .../ContinuedFractions/Computation/Basic.lean | 9 +- .../Computation/TerminatesIffRat.lean | 11 +- .../Computation/Translations.lean | 24 +- .../ContinuedFractions/ConvergentsEquiv.lean | 19 +- .../ContinuedFractions/TerminatedStable.lean | 12 +- .../ContinuedFractions/Translations.lean | 15 +- Mathbin/Algebra/Homology/Additive.lean | 40 +- .../DoldKan/Compatibility.lean | 205 +++++ .../AlgebraicTopology/SimplexCategory.lean | 7 +- Mathbin/All.lean | 30 +- Mathbin/Analysis/Calculus/ContDiff.lean | 4 +- Mathbin/Analysis/Convex/Between.lean | 13 +- Mathbin/Analysis/Convex/Combination.lean | 8 +- Mathbin/CategoryTheory/Category/Twop.lean | 148 ++-- .../CategoryTheory/DifferentialObject.lean | 64 +- Mathbin/CategoryTheory/GradedObject.lean | 33 +- .../Idempotents/HomologicalComplex.lean | 14 +- .../CategoryTheory/Idempotents/Karoubi.lean | 13 +- Mathbin/CategoryTheory/Shift.lean | 557 ------------ Mathbin/CategoryTheory/Shift/Basic.lean | 818 ++++++++++++++++++ .../CategoryTheory/Triangulated/Basic.lean | 38 +- .../Triangulated/Pretriangulated.lean | 96 +- .../CategoryTheory/Triangulated/Rotate.lean | 287 +----- .../SimpleGraph/Connectivity.lean | 131 ++- .../Combinatorics/SimpleGraph/Ends/Defs.lean | 5 +- Mathbin/Data/Finset/Basic.lean | 4 +- Mathbin/Data/Finset/Image.lean | 28 +- Mathbin/Data/Finset/Pointwise.lean | 48 +- Mathbin/Data/Seq/Parallel.lean | 64 +- Mathbin/Data/Seq/Seq.lean | 497 +++++------ Mathbin/Data/Seq/Wseq.lean | 686 +++++++-------- Mathbin/Data/Set/Pointwise/Smul.lean | 4 +- Mathbin/Geometry/Euclidean/Basic.lean | 6 +- Mathbin/Geometry/Euclidean/Circumcenter.lean | 10 +- Mathbin/Geometry/Euclidean/MongePoint.lean | 4 +- Mathbin/Geometry/Manifold/ContMdiff.lean | 30 +- .../Geometry/Manifold/ContMdiffMfderiv.lean | 240 +---- Mathbin/Geometry/Manifold/Mfderiv.lean | 113 ++- Mathbin/Geometry/Manifold/TangentBundle.lean | 767 ---------------- .../Geometry/Manifold/VectorBundle/Basic.lean | 154 +++- .../Manifold/VectorBundle/Tangent.lean | 144 ++- Mathbin/LinearAlgebra/AffineSpace/Basis.lean | 12 +- .../AffineSpace/Combination.lean | 83 +- .../AffineSpace/Independent.lean | 23 +- Mathbin/LinearAlgebra/AffineSpace/Matrix.lean | 4 +- .../Group/FundamentalDomain.lean | 114 ++- Mathbin/Order/Category/BddDistLat.lean | 125 +++ Mathbin/Order/Category/BddLat.lean | 184 ++++ Mathbin/Order/Category/BddOrd.lean | 125 +++ Mathbin/Order/Category/BoolAlg.lean | 31 +- .../Order/Category/BoundedDistribLattice.lean | 127 --- Mathbin/Order/Category/BoundedLattice.lean | 194 ----- Mathbin/Order/Category/BoundedOrder.lean | 128 --- Mathbin/Order/Category/CompleteLat.lean | 107 +++ Mathbin/Order/Category/CompleteLattice.lean | 109 --- Mathbin/Order/Category/DistLat.lean | 99 +++ Mathbin/Order/Category/DistribLattice.lean | 101 --- Mathbin/Order/Category/FinBoolAlg.lean | 14 +- Mathbin/Order/Category/FinPartOrd.lean | 119 +++ Mathbin/Order/Category/FinPartialOrder.lean | 121 --- .../Order/Category/{Frame.lean => Frm.lean} | 50 +- Mathbin/Order/Category/HeytAlg.lean | 10 +- .../Order/Category/{Lattice.lean => Lat.lean} | 65 +- Mathbin/Order/Category/LinOrd.lean | 96 ++ Mathbin/Order/Category/LinearOrder.lean | 98 --- Mathbin/Order/Category/NonemptyFinLinOrd.lean | 18 +- Mathbin/Order/Category/PartOrd.lean | 134 +++ Mathbin/Order/Category/PartialOrder.lean | 137 --- .../Category/{Preorder.lean => Preord.lean} | 55 +- Mathbin/Order/Category/Semilat.lean | 217 +++++ Mathbin/Order/Category/Semilattice.lean | 223 ----- Mathbin/Order/CompleteBooleanAlgebra.lean | 4 +- Mathbin/Order/Hom/CompleteLattice.lean | 8 +- Mathbin/Order/UpperLower/LocallyFinite.lean | 38 + Mathbin/RepresentationTheory/Basic.lean | 8 +- Mathbin/RingTheory/Ideal/Norm.lean | 30 +- .../RingTheory/Ideal/QuotientOperations.lean | 9 +- .../MvPolynomial/WeightedHomogeneous.lean | 481 ++++++++++ Mathbin/SetTheory/Zfc/Basic.lean | 17 +- Mathbin/SetTheory/Zfc/Ordinal.lean | 17 +- Mathbin/Topology/Category/Locale.lean | 8 +- Mathbin/Topology/FiberBundle/Basic.lean | 224 +++-- Mathbin/Topology/VectorBundle/Basic.lean | 46 +- README.md | 2 +- lake-manifest.json | 8 +- lakefile.lean | 4 +- upstream-rev | 2 +- 89 files changed, 4895 insertions(+), 4357 deletions(-) create mode 100644 Mathbin/AlgebraicTopology/DoldKan/Compatibility.lean delete mode 100644 Mathbin/CategoryTheory/Shift.lean create mode 100644 Mathbin/CategoryTheory/Shift/Basic.lean delete mode 100644 Mathbin/Geometry/Manifold/TangentBundle.lean create mode 100644 Mathbin/Order/Category/BddDistLat.lean create mode 100644 Mathbin/Order/Category/BddLat.lean create mode 100644 Mathbin/Order/Category/BddOrd.lean delete mode 100644 Mathbin/Order/Category/BoundedDistribLattice.lean delete mode 100644 Mathbin/Order/Category/BoundedLattice.lean delete mode 100644 Mathbin/Order/Category/BoundedOrder.lean create mode 100644 Mathbin/Order/Category/CompleteLat.lean delete mode 100644 Mathbin/Order/Category/CompleteLattice.lean create mode 100644 Mathbin/Order/Category/DistLat.lean delete mode 100644 Mathbin/Order/Category/DistribLattice.lean create mode 100644 Mathbin/Order/Category/FinPartOrd.lean delete mode 100644 Mathbin/Order/Category/FinPartialOrder.lean rename Mathbin/Order/Category/{Frame.lean => Frm.lean} (68%) rename Mathbin/Order/Category/{Lattice.lean => Lat.lean} (50%) create mode 100644 Mathbin/Order/Category/LinOrd.lean delete mode 100644 Mathbin/Order/Category/LinearOrder.lean create mode 100644 Mathbin/Order/Category/PartOrd.lean delete mode 100644 Mathbin/Order/Category/PartialOrder.lean rename Mathbin/Order/Category/{Preorder.lean => Preord.lean} (60%) create mode 100644 Mathbin/Order/Category/Semilat.lean delete mode 100644 Mathbin/Order/Category/Semilattice.lean create mode 100644 Mathbin/Order/UpperLower/LocallyFinite.lean create mode 100644 Mathbin/RingTheory/MvPolynomial/WeightedHomogeneous.lean diff --git a/Mathbin/Algebra/ContinuedFractions/Basic.lean b/Mathbin/Algebra/ContinuedFractions/Basic.lean index 901710ca0a..c981303c3a 100644 --- a/Mathbin/Algebra/ContinuedFractions/Basic.lean +++ b/Mathbin/Algebra/ContinuedFractions/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann ! This file was ported from Lean 3 source module algebra.continued_fractions.basic -! leanprover-community/mathlib commit 2738d2ca56cbc63be80c3bd48e9ed90ad94e947d +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -60,6 +60,7 @@ protected structure GeneralizedContinuedFraction.Pair where open GeneralizedContinuedFraction +/- ./././Mathport/Syntax/Translate/Command.lean:224:11: unsupported: unusual advanced open style -/ /-! Interlude: define some expected coercions and instances. -/ @@ -114,7 +115,7 @@ For convenience, one often writes `[h; (a₀, b₀), (a₁, b₁), (a₂, b₂), -/ structure GeneralizedContinuedFraction where h : α - s : SeqCat <| Pair α + s : Seq <| Pair α #align generalized_continued_fraction GeneralizedContinuedFraction variable {α} @@ -123,19 +124,19 @@ namespace GeneralizedContinuedFraction /-- Constructs a generalized continued fraction without fractional part. -/ def ofInteger (a : α) : GeneralizedContinuedFraction α := - ⟨a, SeqCat.nil⟩ + ⟨a, Seq.nil⟩ #align generalized_continued_fraction.of_integer GeneralizedContinuedFraction.ofInteger instance [Inhabited α] : Inhabited (GeneralizedContinuedFraction α) := ⟨ofInteger default⟩ /-- Returns the sequence of partial numerators `aᵢ` of `g`. -/ -def partialNumerators (g : GeneralizedContinuedFraction α) : SeqCat α := +def partialNumerators (g : GeneralizedContinuedFraction α) : Seq α := g.s.map Pair.a #align generalized_continued_fraction.partial_numerators GeneralizedContinuedFraction.partialNumerators /-- Returns the sequence of partial denominators `bᵢ` of `g`. -/ -def partialDenominators (g : GeneralizedContinuedFraction α) : SeqCat α := +def partialDenominators (g : GeneralizedContinuedFraction α) : Seq α := g.s.map Pair.b #align generalized_continued_fraction.partial_denominators GeneralizedContinuedFraction.partialDenominators @@ -167,13 +168,13 @@ variable {β : Type _} [Coe α β] /-- Coerce a gcf by elementwise coercion. -/ instance hasCoeToGeneralizedContinuedFraction : Coe (GeneralizedContinuedFraction α) (GeneralizedContinuedFraction β) := - ⟨fun g => ⟨(g.h : β), (g.s.map coe : SeqCat <| Pair β)⟩⟩ + ⟨fun g => ⟨(g.h : β), (g.s.map coe : Seq <| Pair β)⟩⟩ #align generalized_continued_fraction.has_coe_to_generalized_continued_fraction GeneralizedContinuedFraction.hasCoeToGeneralizedContinuedFraction @[simp, norm_cast] theorem coe_to_generalizedContinuedFraction {g : GeneralizedContinuedFraction α} : (↑(g : GeneralizedContinuedFraction α) : GeneralizedContinuedFraction β) = - ⟨(g.h : β), (g.s.map coe : SeqCat <| Pair β)⟩ := + ⟨(g.h : β), (g.s.map coe : Seq <| Pair β)⟩ := rfl #align generalized_continued_fraction.coe_to_generalized_continued_fraction GeneralizedContinuedFraction.coe_to_generalizedContinuedFraction @@ -389,7 +390,7 @@ Returns the approximation of the fraction described by the given sequence up to For example, `convergents'_aux [(1, 2), (3, 4), (5, 6)] 2 = 1 / (2 + 3 / 4)` and `convergents'_aux [(1, 2), (3, 4), (5, 6)] 0 = 0`. -/ -def convergents'Aux : SeqCat (Pair K) → ℕ → K +def convergents'Aux : Seq (Pair K) → ℕ → K | s, 0 => 0 | s, n + 1 => match s.headI with diff --git a/Mathbin/Algebra/ContinuedFractions/Computation/Approximations.lean b/Mathbin/Algebra/ContinuedFractions/Computation/Approximations.lean index d09e1f49f0..24a43b22cc 100644 --- a/Mathbin/Algebra/ContinuedFractions/Computation/Approximations.lean +++ b/Mathbin/Algebra/ContinuedFractions/Computation/Approximations.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann ! This file was ported from Lean 3 source module algebra.continued_fractions.computation.approximations -! leanprover-community/mathlib commit c85b4d15cfdac3abffce2f449e228a0cf0326912 +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -324,9 +324,9 @@ theorem of_denom_mono : (of v).denominators n ≤ (of v).denominators (n + 1) := by let g := of v cases' Decidable.em <| g.partial_denominators.terminated_at n with terminated not_terminated - · have : g.partial_denominators.nth n = none := by rwa [SeqCat.TerminatedAt] at terminated + · have : g.partial_denominators.nth n = none := by rwa [Stream'.Seq.TerminatedAt] at terminated have : g.terminated_at n := - terminated_at_iff_part_denom_none.elim_right (by rwa [SeqCat.TerminatedAt] at terminated) + terminated_at_iff_part_denom_none.elim_right (by rwa [Stream'.Seq.TerminatedAt] at terminated) have : g.denominators (n + 1) = g.denominators n := denominators_stable_of_terminated n.le_succ this rw [this] diff --git a/Mathbin/Algebra/ContinuedFractions/Computation/Basic.lean b/Mathbin/Algebra/ContinuedFractions/Computation/Basic.lean index 88fabe789c..b3d3046922 100644 --- a/Mathbin/Algebra/ContinuedFractions/Computation/Basic.lean +++ b/Mathbin/Algebra/ContinuedFractions/Computation/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann ! This file was ported from Lean 3 source module algebra.continued_fractions.computation.basic -! leanprover-community/mathlib commit 10d887272d1a72b99da88bcb301d1da9d9d33696 +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -168,10 +168,11 @@ This is just an intermediate representation and users should not (need to) direc it. The setup of rewriting/simplification lemmas that make the definitions easy to use is done in `algebra.continued_fractions.computation.translations`. -/ -protected def seq1 (v : K) : Seq1 <| IntFractPair K := +protected def seq1 (v : K) : Stream'.Seq1 <| IntFractPair K := ⟨IntFractPair.of v,--the head - SeqCat.tail-- take the tail of `int_fract_pair.stream` since the first element is already in the - -- head create a sequence from `int_fract_pair.stream` + Stream'.Seq.tail-- take the tail of `int_fract_pair.stream` since the first element is already in + -- the head + -- create a sequence from `int_fract_pair.stream` ⟨IntFractPair.stream v,-- the underlying stream @stream_isSeq _ _ _ v⟩⟩ diff --git a/Mathbin/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean b/Mathbin/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean index 0f61a26941..d82edd6105 100644 --- a/Mathbin/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean +++ b/Mathbin/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann ! This file was ported from Lean 3 source module algebra.continued_fractions.computation.terminates_iff_rat -! leanprover-community/mathlib commit 134625f523e737f650a6ea7f0c82a6177e45e622 +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -36,6 +36,7 @@ rational, continued fraction, termination namespace GeneralizedContinuedFraction +/- ./././Mathport/Syntax/Translate/Command.lean:224:11: unsupported: unusual advanced open style -/ open GeneralizedContinuedFraction (of) variable {K : Type _} [LinearOrderedField K] [FloorRing K] @@ -227,14 +228,14 @@ theorem coe_of_h_rat_eq : (↑((of q).h : ℚ) : K) = (of v).h := theorem coe_of_s_nth_rat_eq : (((of q).s.get? n).map (Pair.map coe) : Option <| Pair K) = (of v).s.get? n := by - simp only [of, int_fract_pair.seq1, SeqCat.map_nth, SeqCat.nth_tail] - simp only [SeqCat.nth] + simp only [of, int_fract_pair.seq1, seq.map_nth, seq.nth_tail] + simp only [seq.nth] rw [← int_fract_pair.coe_stream_rat_eq v_eq_q] rcases succ_nth_stream_eq : int_fract_pair.stream q (n + 1) with (_ | ⟨_, _⟩) <;> simp [Stream'.map, Stream'.nth, succ_nth_stream_eq] #align generalized_continued_fraction.coe_of_s_nth_rat_eq GeneralizedContinuedFraction.coe_of_s_nth_rat_eq -theorem coe_of_s_rat_eq : ((of q).s.map (Pair.map coe) : SeqCat <| Pair K) = (of v).s := +theorem coe_of_s_rat_eq : ((of q).s.map (Pair.map coe) : Seq <| Pair K) = (of v).s := by ext n rw [← coe_of_s_nth_rat_eq v_eq_q] @@ -254,7 +255,7 @@ theorem coe_of_rat_eq : theorem of_terminates_iff_of_rat_terminates {v : K} {q : ℚ} (v_eq_q : v = (q : K)) : (of v).Terminates ↔ (of q).Terminates := by constructor <;> intro h <;> cases' h with n h <;> use n <;> - simp only [SeqCat.TerminatedAt, (coe_of_s_nth_rat_eq v_eq_q n).symm] at h⊢ <;> + simp only [seq.terminated_at, (coe_of_s_nth_rat_eq v_eq_q n).symm] at h⊢ <;> cases (of q).s.get? n <;> trivial #align generalized_continued_fraction.of_terminates_iff_of_rat_terminates GeneralizedContinuedFraction.of_terminates_iff_of_rat_terminates diff --git a/Mathbin/Algebra/ContinuedFractions/Computation/Translations.lean b/Mathbin/Algebra/ContinuedFractions/Computation/Translations.lean index 25d6c61ad9..809e42559d 100644 --- a/Mathbin/Algebra/ContinuedFractions/Computation/Translations.lean +++ b/Mathbin/Algebra/ContinuedFractions/Computation/Translations.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann ! This file was ported from Lean 3 source module algebra.continued_fractions.computation.translations -! leanprover-community/mathlib commit 10d887272d1a72b99da88bcb301d1da9d9d33696 +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -47,6 +47,7 @@ namespace GeneralizedContinuedFraction open GeneralizedContinuedFraction (of) +/- ./././Mathport/Syntax/Translate/Command.lean:224:11: unsupported: unusual advanced open style -/ -- Fix a discrete linear ordered floor field and a value `v`. variable {K : Type _} [LinearOrderedField K] [FloorRing K] {v : K} @@ -217,7 +218,7 @@ theorem of_terminatedAt_iff_intFractPair_seq1_terminatedAt : theorem of_terminatedAt_n_iff_succ_nth_intFractPair_stream_eq_none : (of v).TerminatedAt n ↔ IntFractPair.stream v (n + 1) = none := by - rw [of_terminated_at_iff_int_fract_pair_seq1_terminated_at, SeqCat.TerminatedAt, + rw [of_terminated_at_iff_int_fract_pair_seq1_terminated_at, Stream'.Seq.TerminatedAt, int_fract_pair.nth_seq1_eq_succ_nth_stream] #align generalized_continued_fraction.of_terminated_at_n_iff_succ_nth_int_fract_pair_stream_eq_none GeneralizedContinuedFraction.of_terminatedAt_n_iff_succ_nth_intFractPair_stream_eq_none @@ -240,7 +241,7 @@ theorem IntFractPair.exists_succ_nth_stream_of_gcf_of_nth_eq_some {gp_n : Pair K ∃ ifp, int_fract_pair.stream v (n + 1) = some ifp ∧ pair.mk 1 (ifp.b : K) = gp_n := by unfold of int_fract_pair.seq1 at s_nth_eq - rwa [SeqCat.map_tail, SeqCat.nth_tail, SeqCat.map_nth, Option.map_eq_some'] at s_nth_eq + rwa [seq.map_tail, seq.nth_tail, seq.map_nth, Option.map_eq_some'] at s_nth_eq cases gp_n_eq injection gp_n_eq with _ ifp_b_eq_gp_n_b exists ifp @@ -255,8 +256,8 @@ theorem nth_of_eq_some_of_succ_nth_intFractPair_stream {ifp_succ_n : IntFractPai (of v).s.get? n = some ⟨1, ifp_succ_n.b⟩ := by unfold of int_fract_pair.seq1 - rw [SeqCat.map_tail, SeqCat.nth_tail, SeqCat.map_nth] - simp [SeqCat.nth, stream_succ_nth_eq] + rw [seq.map_tail, seq.nth_tail, seq.map_nth] + simp [seq.nth, stream_succ_nth_eq] #align generalized_continued_fraction.nth_of_eq_some_of_succ_nth_int_fract_pair_stream GeneralizedContinuedFraction.nth_of_eq_some_of_succ_nth_intFractPair_stream /-- Shows how the entries of the sequence of the computed continued fraction can be obtained by the @@ -282,7 +283,7 @@ theorem of_s_head_aux (v : K) : b := p.b }) := by rw [of, int_fract_pair.seq1, of._match_1] - simp only [SeqCat.map_tail, SeqCat.map, SeqCat.tail, SeqCat.head, SeqCat.nth, Stream'.map] + simp only [seq.map_tail, seq.map, seq.tail, seq.head, seq.nth, Stream'.map] rw [← Stream'.nth_succ, Stream'.nth, Option.map] #align generalized_continued_fraction.of_s_head_aux GeneralizedContinuedFraction.of_s_head_aux @@ -299,14 +300,14 @@ variable (K) /-- If `a` is an integer, then the coefficient sequence of its continued fraction is empty. -/ -theorem of_s_of_int (a : ℤ) : (of (a : K)).s = SeqCat.nil := +theorem of_s_of_int (a : ℤ) : (of (a : K)).s = Seq.nil := haveI h : ∀ n, (of (a : K)).s.get? n = none := by intro n induction' n with n ih · rw [of_s_head_aux, stream_succ_of_int, Option.bind] · exact (of (a : K)).s.Prop ih - SeqCat.ext fun n => (h n).trans (SeqCat.nth_nil n).symm + seq.ext fun n => (h n).trans (seq.nth_nil n).symm #align generalized_continued_fraction.of_s_of_int GeneralizedContinuedFraction.of_s_of_int variable {K} (v) @@ -318,8 +319,7 @@ theorem of_s_succ (n : ℕ) : (of v).s.get? (n + 1) = (of (fract v)⁻¹).s.get? by cases' eq_or_ne (fract v) 0 with h h · obtain ⟨a, rfl⟩ : ∃ a : ℤ, v = a := ⟨⌊v⌋, eq_of_sub_eq_zero h⟩ - rw [fract_int_cast, inv_zero, of_s_of_int, ← cast_zero, of_s_of_int, SeqCat.nth_nil, - SeqCat.nth_nil] + rw [fract_int_cast, inv_zero, of_s_of_int, ← cast_zero, of_s_of_int, seq.nth_nil, seq.nth_nil] cases' eq_or_ne ((of (fract v)⁻¹).s.get? n) none with h₁ h₁ · rwa [h₁, ← terminated_at_iff_s_none, @@ -337,7 +337,7 @@ an element `v` of `K` as the coefficient sequence of that of the inverse of the fractional part of `v`. -/ theorem of_s_tail : (of v).s.tail = (of (fract v)⁻¹).s := - SeqCat.ext fun n => SeqCat.nth_tail (of v).s n ▸ of_s_succ v n + Seq.ext fun n => Seq.nth_tail (of v).s n ▸ of_s_succ v n #align generalized_continued_fraction.of_s_tail GeneralizedContinuedFraction.of_s_tail variable (K) (n) @@ -350,7 +350,7 @@ theorem convergents'_of_int (a : ℤ) : (of (a : K)).convergents' n = a := induction' n with n ih · simp only [zeroth_convergent'_eq_h, of_h_eq_floor, floor_int_cast] · rw [convergents', of_h_eq_floor, floor_int_cast, add_right_eq_self] - exact convergents'_aux_succ_none ((of_s_of_int K a).symm ▸ SeqCat.nth_nil 0) _ + exact convergents'_aux_succ_none ((of_s_of_int K a).symm ▸ seq.nth_nil 0) _ #align generalized_continued_fraction.convergents'_of_int GeneralizedContinuedFraction.convergents'_of_int variable {K} (v) diff --git a/Mathbin/Algebra/ContinuedFractions/ConvergentsEquiv.lean b/Mathbin/Algebra/ContinuedFractions/ConvergentsEquiv.lean index b7aa8cd900..53531f3c7e 100644 --- a/Mathbin/Algebra/ContinuedFractions/ConvergentsEquiv.lean +++ b/Mathbin/Algebra/ContinuedFractions/ConvergentsEquiv.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann ! This file was ported from Lean 3 source module algebra.continued_fractions.convergents_equiv -! leanprover-community/mathlib commit 2738d2ca56cbc63be80c3bd48e9ed90ad94e947d +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -73,7 +73,8 @@ variable {K : Type _} {n : ℕ} namespace GeneralizedContinuedFraction -variable {g : GeneralizedContinuedFraction K} {s : SeqCat <| Pair K} +/- ./././Mathport/Syntax/Translate/Command.lean:224:11: unsupported: unusual advanced open style -/ +variable {g : GeneralizedContinuedFraction K} {s : Seq <| Pair K} section Squash @@ -93,10 +94,10 @@ combines `⟨aₙ, bₙ⟩` and `⟨aₙ₊₁, bₙ₊₁⟩` at position `n` t `squash_seq s 0 = [(a₀, bₒ + a₁ / b₁), (a₁, b₁),...]`. If `s.terminated_at (n + 1)`, then `squash_seq s n = s`. -/ -def squashSeq (s : SeqCat <| Pair K) (n : ℕ) : SeqCat (Pair K) := +def squashSeq (s : Seq <| Pair K) (n : ℕ) : Seq (Pair K) := match Prod.mk (s.get? n) (s.get? (n + 1)) with | ⟨some gp_n, some gp_succ_n⟩ => - SeqCat.nats.zipWith + Seq.nats.zipWith (-- return the squashed value at position `n`; otherwise, do nothing. fun n' gp => if n' = n then ⟨gp_n.a, gp_n.b + gp_succ_n.a / gp_succ_n.b⟩ else gp) s @@ -144,7 +145,7 @@ theorem squashSeq_succ_n_tail_eq_squashSeq_tail_n : none => have : squash_seq s (n + 1) = s := squash_seq_eq_self_of_terminated s_succ_succ_nth_eq cases s_succ_nth_eq : s.nth (n + 1) <;> - simp only [squash_seq, SeqCat.nth_tail, s_succ_nth_eq, s_succ_succ_nth_eq] + simp only [squash_seq, seq.nth_tail, s_succ_nth_eq, s_succ_succ_nth_eq] case some => obtain ⟨gp_succ_n, s_succ_nth_eq⟩ : ∃ gp_succ_n, s.nth (n + 1) = some gp_succ_n; @@ -157,7 +158,7 @@ theorem squashSeq_succ_n_tail_eq_squashSeq_tail_n : · have : s.tail.nth m = s.nth (m + 1) := s.nth_tail m cases s_succ_mth_eq : s.nth (m + 1) all_goals have s_tail_mth_eq := this.trans s_succ_mth_eq - · simp only [*, squash_seq, SeqCat.nth_tail, SeqCat.nth_zipWith, Option.map₂_none_right] + · simp only [*, squash_seq, seq.nth_tail, seq.nth_zip_with, Option.map₂_none_right] · simp [*, squash_seq] #align generalized_continued_fraction.squash_seq_succ_n_tail_eq_squash_seq_tail_n GeneralizedContinuedFraction.squashSeq_succ_n_tail_eq_squashSeq_tail_n @@ -177,7 +178,7 @@ theorem succ_succ_nth_convergent'_aux_eq_succ_nth_convergent'_aux_squashSeq : exact s.ge_stable zero_le_one s_succ_nth_eq have : (squash_seq s 0).headI = some ⟨gp_head.a, gp_head.b + gp_succ_n.a / gp_succ_n.b⟩ := squash_seq_nth_of_not_terminated s_head_eq s_succ_nth_eq - simp [*, convergents'_aux, SeqCat.head, SeqCat.nth_tail] + simp [*, convergents'_aux, seq.head, seq.nth_tail] case succ => obtain ⟨gp_head, s_head_eq⟩ : ∃ gp_head, s.head = some gp_head exact s.ge_stable (m + 2).zero_le s_succ_nth_eq @@ -188,7 +189,7 @@ theorem succ_succ_nth_convergent'_aux_eq_succ_nth_convergent'_aux_squashSeq : have : convergents'_aux s.tail (m + 2) = convergents'_aux (squash_seq s.tail m) (m + 1) := by refine' IH gp_succ_n _ - simpa [SeqCat.nth_tail] using s_succ_nth_eq + simpa [seq.nth_tail] using s_succ_nth_eq have : (squash_seq s (m + 1)).headI = some gp_head := (squash_seq_nth_of_lt m.succ_pos).trans s_head_eq simp only [*, convergents'_aux, squash_seq_succ_n_tail_eq_squash_seq_tail_n] @@ -237,7 +238,7 @@ theorem succ_nth_convergent'_eq_squashGcf_nth_convergent' : cases n case zero => cases g_s_head_eq : g.s.nth 0 <;> - simp [g_s_head_eq, squash_gcf, convergents', convergents'_aux, SeqCat.head] + simp [g_s_head_eq, squash_gcf, convergents', convergents'_aux, seq.head] case succ => simp only [succ_succ_nth_convergent'_aux_eq_succ_nth_convergent'_aux_squash_seq, convergents', squash_gcf] diff --git a/Mathbin/Algebra/ContinuedFractions/TerminatedStable.lean b/Mathbin/Algebra/ContinuedFractions/TerminatedStable.lean index c63261b51c..d15a9b5987 100644 --- a/Mathbin/Algebra/ContinuedFractions/TerminatedStable.lean +++ b/Mathbin/Algebra/ContinuedFractions/TerminatedStable.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann ! This file was ported from Lean 3 source module algebra.continued_fractions.terminated_stable -! leanprover-community/mathlib commit 134625f523e737f650a6ea7f0c82a6177e45e622 +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -21,6 +21,7 @@ We show that the continuants and convergents of a gcf stabilise once the gcf ter namespace GeneralizedContinuedFraction +/- ./././Mathport/Syntax/Translate/Command.lean:224:11: unsupported: unusual advanced open style -/ variable {K : Type _} {g : GeneralizedContinuedFraction K} {n m : ℕ} /-- If a gcf terminated at position `n`, it also terminated at `m ≥ n`.-/ @@ -47,23 +48,22 @@ theorem continuantsAux_stable_of_terminated (n_lt_m : n < m) (terminated_at_n : exact terminated_stable (Nat.le_add_right _ _) terminated_at_n #align generalized_continued_fraction.continuants_aux_stable_of_terminated GeneralizedContinuedFraction.continuantsAux_stable_of_terminated -theorem convergents'Aux_stable_step_of_terminated {s : SeqCat <| Pair K} +theorem convergents'Aux_stable_step_of_terminated {s : Seq <| Pair K} (terminated_at_n : s.TerminatedAt n) : convergents'Aux s (n + 1) = convergents'Aux s n := by change s.nth n = none at terminated_at_n induction' n with n IH generalizing s - case zero => simp only [convergents'_aux, terminated_at_n, SeqCat.head] + case zero => simp only [convergents'_aux, terminated_at_n, seq.head] case succ => cases' s_head_eq : s.head with gp_head case none => simp only [convergents'_aux, s_head_eq] case some => - have : s.tail.terminated_at n := by - simp only [SeqCat.TerminatedAt, s.nth_tail, terminated_at_n] + have : s.tail.terminated_at n := by simp only [seq.terminated_at, s.nth_tail, terminated_at_n] simp only [convergents'_aux, s_head_eq, IH this] #align generalized_continued_fraction.convergents'_aux_stable_step_of_terminated GeneralizedContinuedFraction.convergents'Aux_stable_step_of_terminated -theorem convergents'Aux_stable_of_terminated {s : SeqCat <| Pair K} (n_le_m : n ≤ m) +theorem convergents'Aux_stable_of_terminated {s : Seq <| Pair K} (n_le_m : n ≤ m) (terminated_at_n : s.TerminatedAt n) : convergents'Aux s m = convergents'Aux s n := by induction' n_le_m with m n_le_m IH diff --git a/Mathbin/Algebra/ContinuedFractions/Translations.lean b/Mathbin/Algebra/ContinuedFractions/Translations.lean index 1b3cee9bff..a091bebd8a 100644 --- a/Mathbin/Algebra/ContinuedFractions/Translations.lean +++ b/Mathbin/Algebra/ContinuedFractions/Translations.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann ! This file was ported from Lean 3 source module algebra.continued_fractions.translations -! leanprover-community/mathlib commit 10d887272d1a72b99da88bcb301d1da9d9d33696 +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -22,6 +22,7 @@ Some simple translation lemmas between the different definitions of functions de namespace GeneralizedContinuedFraction +/- ./././Mathport/Syntax/Translate/Command.lean:224:11: unsupported: unusual advanced open style -/ section General /-! @@ -66,13 +67,13 @@ theorem part_denom_eq_s_b {gp : Pair α} (s_nth_eq : g.s.get? n = some gp) : theorem exists_s_a_of_part_num {a : α} (nth_part_num_eq : g.partialNumerators.get? n = some a) : ∃ gp, g.s.get? n = some gp ∧ gp.a = a := by - simpa [partial_numerators, SeqCat.map_nth] using nth_part_num_eq + simpa [partial_numerators, seq.map_nth] using nth_part_num_eq #align generalized_continued_fraction.exists_s_a_of_part_num GeneralizedContinuedFraction.exists_s_a_of_part_num theorem exists_s_b_of_part_denom {b : α} (nth_part_denom_eq : g.partialDenominators.get? n = some b) : ∃ gp, g.s.get? n = some gp ∧ gp.b = b := by - simpa [partial_denominators, SeqCat.map_nth] using nth_part_denom_eq + simpa [partial_denominators, seq.map_nth] using nth_part_denom_eq #align generalized_continued_fraction.exists_s_b_of_part_denom GeneralizedContinuedFraction.exists_s_b_of_part_denom end General @@ -167,7 +168,7 @@ theorem first_denominator_eq {gp : Pair K} (zeroth_s_eq : g.s.get? 0 = some gp) #align generalized_continued_fraction.first_denominator_eq GeneralizedContinuedFraction.first_denominator_eq @[simp] -theorem zeroth_convergent'_aux_eq_zero {s : SeqCat <| Pair K} : convergents'Aux s 0 = 0 := +theorem zeroth_convergent'_aux_eq_zero {s : Seq <| Pair K} : convergents'Aux s 0 = 0 := rfl #align generalized_continued_fraction.zeroth_convergent'_aux_eq_zero GeneralizedContinuedFraction.zeroth_convergent'_aux_eq_zero @@ -175,12 +176,12 @@ theorem zeroth_convergent'_aux_eq_zero {s : SeqCat <| Pair K} : convergents'Aux theorem zeroth_convergent'_eq_h : g.convergents' 0 = g.h := by simp [convergents'] #align generalized_continued_fraction.zeroth_convergent'_eq_h GeneralizedContinuedFraction.zeroth_convergent'_eq_h -theorem convergents'Aux_succ_none {s : SeqCat (Pair K)} (h : s.headI = none) (n : ℕ) : +theorem convergents'Aux_succ_none {s : Seq (Pair K)} (h : s.headI = none) (n : ℕ) : convergents'Aux s (n + 1) = 0 := by rw [convergents'_aux, h, convergents'_aux._match_1] #align generalized_continued_fraction.convergents'_aux_succ_none GeneralizedContinuedFraction.convergents'Aux_succ_none -theorem convergents'Aux_succ_some {s : SeqCat (Pair K)} {p : Pair K} (h : s.headI = some p) - (n : ℕ) : convergents'Aux s (n + 1) = p.a / (p.b + convergents'Aux s.tail n) := by +theorem convergents'Aux_succ_some {s : Seq (Pair K)} {p : Pair K} (h : s.headI = some p) (n : ℕ) : + convergents'Aux s (n + 1) = p.a / (p.b + convergents'Aux s.tail n) := by rw [convergents'_aux, h, convergents'_aux._match_1] #align generalized_continued_fraction.convergents'_aux_succ_some GeneralizedContinuedFraction.convergents'Aux_succ_some diff --git a/Mathbin/Algebra/Homology/Additive.lean b/Mathbin/Algebra/Homology/Additive.lean index 464252bc4b..0fdbcec00d 100644 --- a/Mathbin/Algebra/Homology/Additive.lean +++ b/Mathbin/Algebra/Homology/Additive.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison ! This file was ported from Lean 3 source module algebra.homology.additive -! leanprover-community/mathlib commit 88bca0ce5d22ebfd9e73e682e51d60ea13b48347 +! leanprover-community/mathlib commit 200eda15d8ff5669854ff6bcc10aaf37cb70498f ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -157,6 +157,17 @@ def Functor.mapHomologicalComplex (F : V ⥤ W) [F.Additive] (c : ComplexShape rw [← F.map_comp, ← F.map_comp, f.comm] } #align category_theory.functor.map_homological_complex CategoryTheory.Functor.mapHomologicalComplex +variable (V) + +/-- The functor on homological complexes induced by the identity functor is +isomorphic to the identity functor. -/ +@[simps] +def Functor.mapHomologicalComplexIdIso (c : ComplexShape ι) : (𝟭 V).mapHomologicalComplex c ≅ 𝟭 _ := + NatIso.ofComponents (fun K => Hom.isoOfComponents (fun i => Iso.refl _) (by tidy)) (by tidy) +#align category_theory.functor.map_homological_complex_id_iso CategoryTheory.Functor.mapHomologicalComplexIdIso + +variable {V} + instance Functor.map_homogical_complex_additive (F : V ⥤ W) [F.Additive] (c : ComplexShape ι) : (F.mapHomologicalComplex c).Additive where #align category_theory.functor.map_homogical_complex_additive CategoryTheory.Functor.map_homogical_complex_additive @@ -203,6 +214,33 @@ theorem NatTrans.mapHomologicalComplex_naturality {c : ComplexShape ι} {F G : V by tidy #align category_theory.nat_trans.map_homological_complex_naturality CategoryTheory.NatTrans.mapHomologicalComplex_naturality +/-- A natural isomorphism between functors induces a natural isomorphism +between those functors applied to homological complexes. +-/ +@[simps] +def NatIso.mapHomologicalComplex {F G : V ⥤ W} [F.Additive] [G.Additive] (α : F ≅ G) + (c : ComplexShape ι) : F.mapHomologicalComplex c ≅ G.mapHomologicalComplex c + where + Hom := α.Hom.mapHomologicalComplex c + inv := α.inv.mapHomologicalComplex c + hom_inv_id' := by simpa only [← nat_trans.map_homological_complex_comp, α.hom_inv_id] + inv_hom_id' := by simpa only [← nat_trans.map_homological_complex_comp, α.inv_hom_id] +#align category_theory.nat_iso.map_homological_complex CategoryTheory.NatIso.mapHomologicalComplex + +/-- An equivalence of categories induces an equivalences between the respective categories +of homological complex. +-/ +@[simps] +def Equivalence.mapHomologicalComplex (e : V ≌ W) [e.Functor.Additive] (c : ComplexShape ι) : + HomologicalComplex V c ≌ HomologicalComplex W c + where + Functor := e.Functor.mapHomologicalComplex c + inverse := e.inverse.mapHomologicalComplex c + unitIso := + (Functor.mapHomologicalComplexIdIso V c).symm ≪≫ NatIso.mapHomologicalComplex e.unitIso c + counitIso := NatIso.mapHomologicalComplex e.counitIso c ≪≫ Functor.mapHomologicalComplexIdIso W c +#align category_theory.equivalence.map_homological_complex CategoryTheory.Equivalence.mapHomologicalComplex + end CategoryTheory namespace ChainComplex diff --git a/Mathbin/AlgebraicTopology/DoldKan/Compatibility.lean b/Mathbin/AlgebraicTopology/DoldKan/Compatibility.lean new file mode 100644 index 0000000000..7f335711e7 --- /dev/null +++ b/Mathbin/AlgebraicTopology/DoldKan/Compatibility.lean @@ -0,0 +1,205 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou + +! This file was ported from Lean 3 source module algebraic_topology.dold_kan.compatibility +! leanprover-community/mathlib commit 160f568dcf772b2477791c844fc605f2f91f73d1 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Equivalence + +/-! Tools for compatibilities between Dold-Kan equivalences + +The purpose of this file is to introduce tools which will enable the +construction of the Dold-Kan equivalence `simplicial_object C ≌ chain_complex C ℕ` +for a pseudoabelian category `C` from the equivalence +`karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)` and the two +equivalences `simplicial_object C ≅ karoubi (simplicial_object C)` and +`chain_complex C ℕ ≅ karoubi (chain_complex C ℕ)`. + +It is certainly possible to get an equivalence `simplicial_object C ≌ chain_complex C ℕ` +using a compositions of the three equivalences above, but then neither the functor +nor the inverse would have good definitional properties. For example, it would be better +if the inverse functor of the equivalence was exactly the functor +`Γ₀ : simplicial_object C ⥤ chain_complex C ℕ` which was constructed in `functor_gamma.lean`. + +In this file, given four categories `A`, `A'`, `B`, `B'`, equivalences `eA : A ≅ A'`, +`eB : B ≅ B'`, `e' : A' ≅ B'`, functors `F : A ⥤ B'`, `G : B ⥤ A` equipped with certain +compatibilities, we construct successive equivalences: +- `equivalence₀` from `A` to `B'`, which is the composition of `eA` and `e'`. +- `equivalence₁` from `A` to `B'`, with the same inverse functor as `equivalence₀`, +but whose functor is `F`. +- `equivalence₂` from `A` to `B`, which is the composition of `equivalence₁` and the +inverse of `eB`: +- `equivalence` from `A` to `B`, which has the same functor `F ⋙ eB.inverse` as `equivalence₂`, +but whose inverse functor is `G`. + +When extra assumptions are given, we shall also provide simplification lemmas for the +unit and counit isomorphisms of `equivalence`. (TODO) + +-/ + + +open CategoryTheory CategoryTheory.Category + +namespace AlgebraicTopology + +namespace DoldKan + +namespace Compatibility + +variable {A A' B B' : Type _} [Category A] [Category A'] [Category B] [Category B'] (eA : A ≌ A') + (eB : B ≌ B') (e' : A' ≌ B') {F : A ⥤ B'} (hF : eA.Functor ⋙ e'.Functor ≅ F) {G : B ⥤ A} + (hG : eB.Functor ⋙ e'.inverse ≅ G ⋙ eA.Functor) + +/-- A basic equivalence `A ≅ B'` obtained by composing `eA : A ≅ A'` and `e' : A' ≅ B'`. -/ +@[simps Functor inverse unit_iso_hom_app] +def equivalence₀ : A ≌ B' := + eA.trans e' +#align algebraic_topology.dold_kan.compatibility.equivalence₀ AlgebraicTopology.DoldKan.Compatibility.equivalence₀ + +include hF + +variable {eA} {e'} + +/-- An intermediate equivalence `A ≅ B'` whose functor is `F` and whose inverse is +`e'.inverse ⋙ eA.inverse`. -/ +@[simps Functor] +def equivalence₁ : A ≌ B' := + letI : is_equivalence F := + is_equivalence.of_iso hF (is_equivalence.of_equivalence (equivalence₀ eA e')) + F.as_equivalence +#align algebraic_topology.dold_kan.compatibility.equivalence₁ AlgebraicTopology.DoldKan.Compatibility.equivalence₁ + +theorem equivalence₁_inverse : (equivalence₁ hF).inverse = e'.inverse ⋙ eA.inverse := + rfl +#align algebraic_topology.dold_kan.compatibility.equivalence₁_inverse AlgebraicTopology.DoldKan.Compatibility.equivalence₁_inverse + +/-- The counit isomorphism of the equivalence `equivalence₁` between `A` and `B'`. -/ +@[simps] +def equivalence₁CounitIso : (e'.inverse ⋙ eA.inverse) ⋙ F ≅ 𝟭 B' := + calc + (e'.inverse ⋙ eA.inverse) ⋙ F ≅ (e'.inverse ⋙ eA.inverse) ⋙ eA.Functor ⋙ e'.Functor := + isoWhiskerLeft _ hF.symm + _ ≅ e'.inverse ⋙ (eA.inverse ⋙ eA.Functor) ⋙ e'.Functor := (Iso.refl _) + _ ≅ e'.inverse ⋙ 𝟭 _ ⋙ e'.Functor := (isoWhiskerLeft _ (isoWhiskerRight eA.counitIso _)) + _ ≅ e'.inverse ⋙ e'.Functor := (Iso.refl _) + _ ≅ 𝟭 B' := e'.counitIso + +#align algebraic_topology.dold_kan.compatibility.equivalence₁_counit_iso AlgebraicTopology.DoldKan.Compatibility.equivalence₁CounitIso + +theorem equivalence₁CounitIso_eq : (equivalence₁ hF).counitIso = equivalence₁CounitIso hF := + by + ext Y + dsimp [equivalence₀, equivalence₁, is_equivalence.inverse, is_equivalence.of_equivalence] + simp only [equivalence₁_counit_iso_hom_app, CategoryTheory.Functor.map_id, comp_id] +#align algebraic_topology.dold_kan.compatibility.equivalence₁_counit_iso_eq AlgebraicTopology.DoldKan.Compatibility.equivalence₁CounitIso_eq + +/-- The unit isomorphism of the equivalence `equivalence₁` between `A` and `B'`. -/ +@[simps] +def equivalence₁UnitIso : 𝟭 A ≅ F ⋙ e'.inverse ⋙ eA.inverse := + calc + 𝟭 A ≅ eA.Functor ⋙ eA.inverse := eA.unitIso + _ ≅ eA.Functor ⋙ 𝟭 A' ⋙ eA.inverse := (Iso.refl _) + _ ≅ eA.Functor ⋙ (e'.Functor ⋙ e'.inverse) ⋙ eA.inverse := + (isoWhiskerLeft _ (isoWhiskerRight e'.unitIso _)) + _ ≅ (eA.Functor ⋙ e'.Functor) ⋙ e'.inverse ⋙ eA.inverse := (Iso.refl _) + _ ≅ F ⋙ e'.inverse ⋙ eA.inverse := isoWhiskerRight hF _ + +#align algebraic_topology.dold_kan.compatibility.equivalence₁_unit_iso AlgebraicTopology.DoldKan.Compatibility.equivalence₁UnitIso + +theorem equivalence₁UnitIso_eq : (equivalence₁ hF).unitIso = equivalence₁UnitIso hF := + by + ext X + dsimp [equivalence₀, equivalence₁, nat_iso.hcomp, is_equivalence.of_equivalence] + simp only [id_comp, assoc, equivalence₁_unit_iso_hom_app] +#align algebraic_topology.dold_kan.compatibility.equivalence₁_unit_iso_eq AlgebraicTopology.DoldKan.Compatibility.equivalence₁UnitIso_eq + +include eB + +/-- An intermediate equivalence `A ≅ B` obtained as the composition of `equivalence₁` and +the inverse of `eB : B ≌ B'`. -/ +@[simps Functor] +def equivalence₂ : A ≌ B := + (equivalence₁ hF).trans eB.symm +#align algebraic_topology.dold_kan.compatibility.equivalence₂ AlgebraicTopology.DoldKan.Compatibility.equivalence₂ + +theorem equivalence₂_inverse : + (equivalence₂ eB hF).inverse = eB.Functor ⋙ e'.inverse ⋙ eA.inverse := + rfl +#align algebraic_topology.dold_kan.compatibility.equivalence₂_inverse AlgebraicTopology.DoldKan.Compatibility.equivalence₂_inverse + +/-- The counit isomorphism of the equivalence `equivalence₂` between `A` and `B`. -/ +@[simps] +def equivalence₂CounitIso : (eB.Functor ⋙ e'.inverse ⋙ eA.inverse) ⋙ F ⋙ eB.inverse ≅ 𝟭 B := + calc + (eB.Functor ⋙ e'.inverse ⋙ eA.inverse) ⋙ F ⋙ eB.inverse ≅ + eB.Functor ⋙ (e'.inverse ⋙ eA.inverse ⋙ F) ⋙ eB.inverse := + Iso.refl _ + _ ≅ eB.Functor ⋙ 𝟭 _ ⋙ eB.inverse := + (isoWhiskerLeft _ (isoWhiskerRight (equivalence₁CounitIso hF) _)) + _ ≅ eB.Functor ⋙ eB.inverse := (Iso.refl _) + _ ≅ 𝟭 B := eB.unitIso.symm + +#align algebraic_topology.dold_kan.compatibility.equivalence₂_counit_iso AlgebraicTopology.DoldKan.Compatibility.equivalence₂CounitIso + +theorem equivalence₂CounitIso_eq : (equivalence₂ eB hF).counitIso = equivalence₂CounitIso eB hF := + by + ext Y' + dsimp [equivalence₂, iso.refl] + simp only [equivalence₁_counit_iso_eq, equivalence₂_counit_iso_hom_app, + equivalence₁_counit_iso_hom_app, functor.map_comp, assoc] +#align algebraic_topology.dold_kan.compatibility.equivalence₂_counit_iso_eq AlgebraicTopology.DoldKan.Compatibility.equivalence₂CounitIso_eq + +/-- The unit isomorphism of the equivalence `equivalence₂` between `A` and `B`. -/ +@[simps] +def equivalence₂UnitIso : 𝟭 A ≅ (F ⋙ eB.inverse) ⋙ eB.Functor ⋙ e'.inverse ⋙ eA.inverse := + calc + 𝟭 A ≅ F ⋙ e'.inverse ⋙ eA.inverse := equivalence₁UnitIso hF + _ ≅ F ⋙ 𝟭 B' ⋙ e'.inverse ⋙ eA.inverse := (Iso.refl _) + _ ≅ F ⋙ (eB.inverse ⋙ eB.Functor) ⋙ e'.inverse ⋙ eA.inverse := + (isoWhiskerLeft _ (isoWhiskerRight eB.counitIso.symm _)) + _ ≅ (F ⋙ eB.inverse) ⋙ eB.Functor ⋙ e'.inverse ⋙ eA.inverse := Iso.refl _ + +#align algebraic_topology.dold_kan.compatibility.equivalence₂_unit_iso AlgebraicTopology.DoldKan.Compatibility.equivalence₂UnitIso + +theorem equivalence₂UnitIso_eq : (equivalence₂ eB hF).unitIso = equivalence₂UnitIso eB hF := + by + ext X + dsimp [equivalence₂] + simpa only [equivalence₂_unit_iso_hom_app, equivalence₁_unit_iso_eq, + equivalence₁_unit_iso_hom_app, assoc, nat_iso.cancel_nat_iso_hom_left] +#align algebraic_topology.dold_kan.compatibility.equivalence₂_unit_iso_eq AlgebraicTopology.DoldKan.Compatibility.equivalence₂UnitIso_eq + +variable {eB} + +include hG + +/-- The equivalence `A ≅ B` whose functor is `F ⋙ eB.inverse` and +whose inverse is `G : B ≅ A`. -/ +@[simps inverse] +def equivalence : A ≌ B := + letI : is_equivalence G := + by + refine' is_equivalence.of_iso _ (is_equivalence.of_equivalence (equivalence₂ eB hF).symm) + calc + eB.functor ⋙ e'.inverse ⋙ eA.inverse ≅ (eB.functor ⋙ e'.inverse) ⋙ eA.inverse := iso.refl _ + _ ≅ (G ⋙ eA.functor) ⋙ eA.inverse := (iso_whisker_right hG _) + _ ≅ G ⋙ 𝟭 A := (iso_whisker_left _ eA.unit_iso.symm) + _ ≅ G := functor.right_unitor G + + G.as_equivalence.symm +#align algebraic_topology.dold_kan.compatibility.equivalence AlgebraicTopology.DoldKan.Compatibility.equivalence + +theorem equivalence_functor : (equivalence hF hG).Functor = F ⋙ eB.inverse := + rfl +#align algebraic_topology.dold_kan.compatibility.equivalence_functor AlgebraicTopology.DoldKan.Compatibility.equivalence_functor + +end Compatibility + +end DoldKan + +end AlgebraicTopology + diff --git a/Mathbin/AlgebraicTopology/SimplexCategory.lean b/Mathbin/AlgebraicTopology/SimplexCategory.lean index a483ec2552..28c28fd4ca 100644 --- a/Mathbin/AlgebraicTopology/SimplexCategory.lean +++ b/Mathbin/AlgebraicTopology/SimplexCategory.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Scott Morrison, Adam Topaz ! This file was ported from Lean 3 source module algebraic_topology.simplex_category -! leanprover-community/mathlib commit dd1f8496baa505636a82748e6b652165ea888733 +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -932,9 +932,8 @@ to the category attached to the ordered set `{0, 1, ..., n}` -/ @[simps obj map] def toCat : SimplexCategory ⥤ Cat.{0} := SimplexCategory.skeletalFunctor ⋙ - forget₂ NonemptyFinLinOrdCat LinearOrderCat ⋙ - forget₂ LinearOrderCat LatticeCat ⋙ - forget₂ LatticeCat PartialOrderCat ⋙ forget₂ PartialOrderCat PreorderCat ⋙ preorderToCat + forget₂ NonemptyFinLinOrdCat LinOrd ⋙ + forget₂ LinOrd Lat ⋙ forget₂ Lat PartOrd ⋙ forget₂ PartOrd Preord ⋙ preordToCat #align simplex_category.to_Cat SimplexCategory.toCat end SimplexCategory diff --git a/Mathbin/All.lean b/Mathbin/All.lean index fc9c4a49d6..0c36677845 100644 --- a/Mathbin/All.lean +++ b/Mathbin/All.lean @@ -448,6 +448,7 @@ import Mathbin.AlgebraicGeometry.StructureSheaf import Mathbin.AlgebraicTopology.MooreComplex import Mathbin.AlgebraicTopology.AlternatingFaceMapComplex import Mathbin.AlgebraicTopology.CechNerve +import Mathbin.AlgebraicTopology.DoldKan.Compatibility import Mathbin.AlgebraicTopology.DoldKan.Decomposition import Mathbin.AlgebraicTopology.DoldKan.Degeneracies import Mathbin.AlgebraicTopology.DoldKan.EquivalenceAdditive @@ -1054,7 +1055,7 @@ import Mathbin.CategoryTheory.Products.Basic import Mathbin.CategoryTheory.Products.Bifunctor import Mathbin.CategoryTheory.Punit import Mathbin.CategoryTheory.Quotient -import Mathbin.CategoryTheory.Shift +import Mathbin.CategoryTheory.Shift.Basic import Mathbin.CategoryTheory.Sigma.Basic import Mathbin.CategoryTheory.Simple import Mathbin.CategoryTheory.SingleObj @@ -1806,7 +1807,6 @@ import Mathbin.Geometry.Manifold.Metrizable import Mathbin.Geometry.Manifold.Mfderiv import Mathbin.Geometry.Manifold.PartitionOfUnity import Mathbin.Geometry.Manifold.SmoothManifoldWithCorners -import Mathbin.Geometry.Manifold.TangentBundle import Mathbin.Geometry.Manifold.VectorBundle.Basic import Mathbin.Geometry.Manifold.VectorBundle.FiberwiseLinear import Mathbin.Geometry.Manifold.VectorBundle.Pullback @@ -2290,22 +2290,22 @@ import Mathbin.Order.Bounded import Mathbin.Order.BoundedOrder import Mathbin.Order.Bounds.Basic import Mathbin.Order.Bounds.OrderIso +import Mathbin.Order.Category.BddDistLat +import Mathbin.Order.Category.BddLat +import Mathbin.Order.Category.BddOrd import Mathbin.Order.Category.BoolAlg -import Mathbin.Order.Category.BoundedDistribLattice -import Mathbin.Order.Category.BoundedLattice -import Mathbin.Order.Category.BoundedOrder -import Mathbin.Order.Category.CompleteLattice -import Mathbin.Order.Category.DistribLattice +import Mathbin.Order.Category.CompleteLat +import Mathbin.Order.Category.DistLat import Mathbin.Order.Category.FinBoolAlg -import Mathbin.Order.Category.FinPartialOrder -import Mathbin.Order.Category.Frame +import Mathbin.Order.Category.FinPartOrd +import Mathbin.Order.Category.Frm import Mathbin.Order.Category.HeytAlg -import Mathbin.Order.Category.Lattice -import Mathbin.Order.Category.LinearOrder +import Mathbin.Order.Category.Lat +import Mathbin.Order.Category.LinOrd import Mathbin.Order.Category.NonemptyFinLinOrd -import Mathbin.Order.Category.PartialOrder -import Mathbin.Order.Category.Preorder -import Mathbin.Order.Category.Semilattice +import Mathbin.Order.Category.PartOrd +import Mathbin.Order.Category.Preord +import Mathbin.Order.Category.Semilat import Mathbin.Order.Category.OmegaCompletePartialOrder import Mathbin.Order.Chain import Mathbin.Order.Circular @@ -2407,6 +2407,7 @@ import Mathbin.Order.SymmDiff import Mathbin.Order.Synonym import Mathbin.Order.UpperLower.Basic import Mathbin.Order.UpperLower.Hom +import Mathbin.Order.UpperLower.LocallyFinite import Mathbin.Order.WellFounded import Mathbin.Order.WellFoundedSet import Mathbin.Order.WithBot @@ -2542,6 +2543,7 @@ import Mathbin.RingTheory.MvPolynomial.Homogeneous import Mathbin.RingTheory.MvPolynomial.Ideal import Mathbin.RingTheory.MvPolynomial.Symmetric import Mathbin.RingTheory.MvPolynomial.Tower +import Mathbin.RingTheory.MvPolynomial.WeightedHomogeneous import Mathbin.RingTheory.Nakayama import Mathbin.RingTheory.Nilpotent import Mathbin.RingTheory.Noetherian diff --git a/Mathbin/Analysis/Calculus/ContDiff.lean b/Mathbin/Analysis/Calculus/ContDiff.lean index 536de4e284..f0a9655efb 100644 --- a/Mathbin/Analysis/Calculus/ContDiff.lean +++ b/Mathbin/Analysis/Calculus/ContDiff.lean @@ -1,10 +1,10 @@ /- Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Floris van Doorn ! This file was ported from Lean 3 source module analysis.calculus.cont_diff -! leanprover-community/mathlib commit dd6388c44e6f6b4547070b887c5905d5cfe6c9f8 +! leanprover-community/mathlib commit 0187644979f2d3e10a06e916a869c994facd9a87 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ diff --git a/Mathbin/Analysis/Convex/Between.lean b/Mathbin/Analysis/Convex/Between.lean index 898e6769e8..002e0813f0 100644 --- a/Mathbin/Analysis/Convex/Between.lean +++ b/Mathbin/Analysis/Convex/Between.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers ! This file was ported from Lean 3 source module analysis.convex.between -! leanprover-community/mathlib commit 78261225eb5cedc61c5c74ecb44e5b385d13b733 +! leanprover-community/mathlib commit 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -541,18 +541,13 @@ theorem Sbtw.trans_wbtw_right_ne [NoZeroSMulDivisors R V] {w x y z : P} (h₁ : h₁.Wbtw.trans_right_ne h₂ h₁.left_ne #align sbtw.trans_wbtw_right_ne Sbtw.trans_wbtw_right_ne -/- Calls to `affine_combination` are slow to elaborate (generally, not just for this lemma), and -without the use of `@finset.affine_combination R V _ _ _ _ _ _` for at least three of the six -calls in this lemma statement, elaboration of the statement times out (even if the proof is -replaced by `sorry`). -/ theorem Sbtw.affineCombination_of_mem_affineSpan_pair [NoZeroDivisors R] [NoZeroSMulDivisors R V] {ι : Type _} {p : ι → P} (ha : AffineIndependent R p) {w w₁ w₂ : ι → R} {s : Finset ι} (hw : (∑ i in s, w i) = 1) (hw₁ : (∑ i in s, w₁ i) = 1) (hw₂ : (∑ i in s, w₂ i) = 1) - (h : s.affineCombination p w ∈ line[R, s.affineCombination p w₁, s.affineCombination p w₂]) + (h : + s.affineCombination R p w ∈ line[R, s.affineCombination R p w₁, s.affineCombination R p w₂]) {i : ι} (his : i ∈ s) (hs : Sbtw R (w₁ i) (w i) (w₂ i)) : - Sbtw R (@Finset.affineCombination R V _ _ _ _ _ _ s p w₁) - (@Finset.affineCombination R V _ _ _ _ _ _ s p w) - (@Finset.affineCombination R V _ _ _ _ _ _ s p w₂) := + Sbtw R (s.affineCombination R p w₁) (s.affineCombination R p w) (s.affineCombination R p w₂) := by rw [affineCombination_mem_affineSpan_pair ha hw hw₁ hw₂] at h rcases h with ⟨r, hr⟩ diff --git a/Mathbin/Analysis/Convex/Combination.lean b/Mathbin/Analysis/Convex/Combination.lean index d4a8c85a3c..7191a26b11 100644 --- a/Mathbin/Analysis/Convex/Combination.lean +++ b/Mathbin/Analysis/Convex/Combination.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudriashov ! This file was ported from Lean 3 source module analysis.convex.combination -! leanprover-community/mathlib commit 2f4cdce0c2f2f3b8cd58f05d556d03b468e1eb2e +! leanprover-community/mathlib commit 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -253,7 +253,7 @@ theorem Finset.centerMass_id_mem_convexHull (t : Finset E) {w : E → R} (hw₀ #align finset.center_mass_id_mem_convex_hull Finset.centerMass_id_mem_convexHull theorem affineCombination_eq_centerMass {ι : Type _} {t : Finset ι} {p : ι → E} {w : ι → R} - (hw₂ : (∑ i in t, w i) = 1) : affineCombination t p w = centerMass t w p := + (hw₂ : (∑ i in t, w i) = 1) : t.affineCombination R p w = centerMass t w p := by rw [affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one _ w _ hw₂ (0 : E), Finset.weightedVsubOfPoint_apply, vadd_eq_add, add_zero, t.center_mass_eq_of_sum_1 _ hw₂] @@ -262,7 +262,7 @@ theorem affineCombination_eq_centerMass {ι : Type _} {t : Finset ι} {p : ι theorem affineCombination_mem_convexHull {s : Finset ι} {v : ι → E} {w : ι → R} (hw₀ : ∀ i ∈ s, 0 ≤ w i) (hw₁ : s.Sum w = 1) : - s.affineCombination v w ∈ convexHull R (range v) := + s.affineCombination R v w ∈ convexHull R (range v) := by rw [affineCombination_eq_centerMass hw₁] apply s.center_mass_mem_convex_hull hw₀ @@ -292,7 +292,7 @@ theorem convexHull_range_eq_exists_affineCombination (v : ι → E) : convexHull R (range v) = { x | ∃ (s : Finset ι)(w : ι → R)(hw₀ : ∀ i ∈ s, 0 ≤ w i)(hw₁ : s.Sum w = 1), - s.affineCombination v w = x } := + s.affineCombination R v w = x } := by refine' subset.antisymm (convexHull_min _ _) _ · intro x hx diff --git a/Mathbin/CategoryTheory/Category/Twop.lean b/Mathbin/CategoryTheory/Category/Twop.lean index 791fa4f13f..98de0246c1 100644 --- a/Mathbin/CategoryTheory/Category/Twop.lean +++ b/Mathbin/CategoryTheory/Category/Twop.lean @@ -29,68 +29,92 @@ universe u variable {α β : Type _} +#print TwoP /- /-- The category of two-pointed types. -/ -structure Twop : Type (u + 1) where +structure TwoP : Type (u + 1) where pt : Type u toTwoPointing : TwoPointing X -#align Twop Twop +#align Twop TwoP +-/ -namespace Twop +namespace TwoP -instance : CoeSort Twop (Type _) := +instance : CoeSort TwoP (Type _) := ⟨X⟩ -attribute [protected] Twop.X +attribute [protected] TwoP.X +#print TwoP.of /- /-- Turns a two-pointing into a two-pointed type. -/ -def of {X : Type _} (to_two_pointing : TwoPointing X) : Twop := +def of {X : Type _} (to_two_pointing : TwoPointing X) : TwoP := ⟨X, to_two_pointing⟩ -#align Twop.of Twop.of +#align Twop.of TwoP.of +-/ +#print TwoP.coe_of /- @[simp] theorem coe_of {X : Type _} (to_two_pointing : TwoPointing X) : ↥(of to_two_pointing) = X := rfl -#align Twop.coe_of Twop.coe_of +#align Twop.coe_of TwoP.coe_of +-/ alias of ← _root_.two_pointing.Twop -#align two_pointing.Twop TwoPointing.twop +#align two_pointing.Twop TwoPointing.TwoP -instance : Inhabited Twop := +instance : Inhabited TwoP := ⟨of TwoPointing.bool⟩ +#print TwoP.toBipointed /- /-- Turns a two-pointed type into a bipointed type, by forgetting that the pointed elements are distinct. -/ -def toBipointed (X : Twop) : Bipointed := +def toBipointed (X : TwoP) : Bipointed := X.toTwoPointing.toProd.Bipointed -#align Twop.to_Bipointed Twop.toBipointed +#align Twop.to_Bipointed TwoP.toBipointed +-/ +#print TwoP.coe_toBipointed /- @[simp] -theorem coe_toBipointed (X : Twop) : ↥X.toBipointed = ↥X := +theorem coe_toBipointed (X : TwoP) : ↥X.toBipointed = ↥X := rfl -#align Twop.coe_to_Bipointed Twop.coe_toBipointed +#align Twop.coe_to_Bipointed TwoP.coe_toBipointed +-/ -instance largeCategory : LargeCategory Twop := +#print TwoP.largeCategory /- +instance largeCategory : LargeCategory TwoP := InducedCategory.category toBipointed -#align Twop.large_category Twop.largeCategory +#align Twop.large_category TwoP.largeCategory +-/ -instance concreteCategory : ConcreteCategory Twop := +#print TwoP.concreteCategory /- +instance concreteCategory : ConcreteCategory TwoP := InducedCategory.concreteCategory toBipointed -#align Twop.concrete_category Twop.concreteCategory +#align Twop.concrete_category TwoP.concreteCategory +-/ -instance hasForgetToBipointed : HasForget₂ Twop Bipointed := +#print TwoP.hasForgetToBipointed /- +instance hasForgetToBipointed : HasForget₂ TwoP Bipointed := InducedCategory.hasForget₂ toBipointed -#align Twop.has_forget_to_Bipointed Twop.hasForgetToBipointed +#align Twop.has_forget_to_Bipointed TwoP.hasForgetToBipointed +-/ +#print TwoP.swap /- /-- Swaps the pointed elements of a two-pointed type. `two_pointing.swap` as a functor. -/ @[simps] -def swap : Twop ⥤ Twop where +def swap : TwoP ⥤ TwoP where obj X := ⟨X, X.toTwoPointing.symm⟩ map X Y f := ⟨f.toFun, f.map_snd, f.map_fst⟩ -#align Twop.swap Twop.swap +#align Twop.swap TwoP.swap +-/ +/- warning: Twop.swap_equiv -> TwoP.swapEquiv is a dubious translation: +lean 3 declaration is + CategoryTheory.Equivalence.{u1, u1, succ u1, succ u1} TwoP.{u1} TwoP.largeCategory.{u1} TwoP.{u1} TwoP.largeCategory.{u1} +but is expected to have type + CategoryTheory.Equivalence.{u1, u1, succ u1, succ u1} TwoP.{u1} TwoP.{u1} TwoP.largeCategory.{u1} TwoP.largeCategory.{u1} +Case conversion may be inaccurate. Consider using '#align Twop.swap_equiv TwoP.swapEquivₓ'. -/ /-- The equivalence between `Twop` and itself induced by `prod.swap` both ways. -/ @[simps] -def swapEquiv : Twop ≌ Twop := +def swapEquiv : TwoP ≌ TwoP := Equivalence.mk swap swap (NatIso.ofComponents (fun X => @@ -102,66 +126,87 @@ def swapEquiv : Twop ≌ Twop := { Hom := ⟨id, rfl, rfl⟩ inv := ⟨id, rfl, rfl⟩ }) fun X Y f => rfl) -#align Twop.swap_equiv Twop.swapEquiv - +#align Twop.swap_equiv TwoP.swapEquiv + +/- warning: Twop.swap_equiv_symm -> TwoP.swapEquiv_symm is a dubious translation: +lean 3 declaration is + Eq.{succ (succ u1)} (CategoryTheory.Equivalence.{u1, u1, succ u1, succ u1} TwoP.{u1} TwoP.largeCategory.{u1} TwoP.{u1} TwoP.largeCategory.{u1}) (CategoryTheory.Equivalence.symm.{u1, u1, succ u1, succ u1} TwoP.{u1} TwoP.largeCategory.{u1} TwoP.{u1} TwoP.largeCategory.{u1} TwoP.swapEquiv.{u1}) TwoP.swapEquiv.{u1} +but is expected to have type + Eq.{succ (succ u1)} (CategoryTheory.Equivalence.{u1, u1, succ u1, succ u1} TwoP.{u1} TwoP.{u1} TwoP.largeCategory.{u1} TwoP.largeCategory.{u1}) (CategoryTheory.Equivalence.symm.{u1, u1, succ u1, succ u1} TwoP.{u1} TwoP.largeCategory.{u1} TwoP.{u1} TwoP.largeCategory.{u1} TwoP.swapEquiv.{u1}) TwoP.swapEquiv.{u1} +Case conversion may be inaccurate. Consider using '#align Twop.swap_equiv_symm TwoP.swapEquiv_symmₓ'. -/ @[simp] theorem swapEquiv_symm : swapEquiv.symm = swapEquiv := rfl -#align Twop.swap_equiv_symm Twop.swapEquiv_symm +#align Twop.swap_equiv_symm TwoP.swapEquiv_symm -end Twop +end TwoP +#print TwoP_swap_comp_forget_to_Bipointed /- @[simp] -theorem twop_swap_comp_forget_to_bipointed : - Twop.swap ⋙ forget₂ Twop Bipointed = forget₂ Twop Bipointed ⋙ Bipointed.swap := +theorem TwoP_swap_comp_forget_to_Bipointed : + TwoP.swap ⋙ forget₂ TwoP Bipointed = forget₂ TwoP Bipointed ⋙ Bipointed.swap := rfl -#align Twop_swap_comp_forget_to_Bipointed twop_swap_comp_forget_to_bipointed +#align Twop_swap_comp_forget_to_Bipointed TwoP_swap_comp_forget_to_Bipointed +-/ +#print pointedToTwoPFst /- /-- The functor from `Pointed` to `Twop` which adds a second point. -/ @[simps] -def pointedToTwopFst : Pointed.{u} ⥤ Twop +def pointedToTwoPFst : Pointed.{u} ⥤ TwoP where obj X := ⟨Option X, ⟨X.point, none⟩, some_ne_none _⟩ map X Y f := ⟨Option.map f.toFun, congr_arg _ f.map_point, rfl⟩ map_id' X := Bipointed.Hom.ext _ _ Option.map_id map_comp' X Y Z f g := Bipointed.Hom.ext _ _ (Option.map_comp_map _ _).symm -#align Pointed_to_Twop_fst pointedToTwopFst +#align Pointed_to_Twop_fst pointedToTwoPFst +-/ +#print pointedToTwoPSnd /- /-- The functor from `Pointed` to `Twop` which adds a first point. -/ @[simps] -def pointedToTwopSnd : Pointed.{u} ⥤ Twop +def pointedToTwoPSnd : Pointed.{u} ⥤ TwoP where obj X := ⟨Option X, ⟨none, X.point⟩, (some_ne_none _).symm⟩ map X Y f := ⟨Option.map f.toFun, rfl, congr_arg _ f.map_point⟩ map_id' X := Bipointed.Hom.ext _ _ Option.map_id map_comp' X Y Z f g := Bipointed.Hom.ext _ _ (Option.map_comp_map _ _).symm -#align Pointed_to_Twop_snd pointedToTwopSnd +#align Pointed_to_Twop_snd pointedToTwoPSnd +-/ +#print pointedToTwoPFst_comp_swap /- @[simp] -theorem pointedToTwopFst_comp_swap : pointedToTwopFst ⋙ Twop.swap = pointedToTwopSnd := +theorem pointedToTwoPFst_comp_swap : pointedToTwoPFst ⋙ TwoP.swap = pointedToTwoPSnd := rfl -#align Pointed_to_Twop_fst_comp_swap pointedToTwopFst_comp_swap +#align Pointed_to_Twop_fst_comp_swap pointedToTwoPFst_comp_swap +-/ +#print pointedToTwoPSnd_comp_swap /- @[simp] -theorem pointedToTwopSnd_comp_swap : pointedToTwopSnd ⋙ Twop.swap = pointedToTwopFst := +theorem pointedToTwoPSnd_comp_swap : pointedToTwoPSnd ⋙ TwoP.swap = pointedToTwoPFst := rfl -#align Pointed_to_Twop_snd_comp_swap pointedToTwopSnd_comp_swap +#align Pointed_to_Twop_snd_comp_swap pointedToTwoPSnd_comp_swap +-/ +#print pointedToTwoPFst_comp_forget_to_bipointed /- @[simp] -theorem pointedToTwopFst_comp_forget_to_bipointed : - pointedToTwopFst ⋙ forget₂ Twop Bipointed = pointedToBipointedFst := +theorem pointedToTwoPFst_comp_forget_to_bipointed : + pointedToTwoPFst ⋙ forget₂ TwoP Bipointed = pointedToBipointedFst := rfl -#align Pointed_to_Twop_fst_comp_forget_to_Bipointed pointedToTwopFst_comp_forget_to_bipointed +#align Pointed_to_Twop_fst_comp_forget_to_Bipointed pointedToTwoPFst_comp_forget_to_bipointed +-/ +#print pointedToTwoPSnd_comp_forget_to_bipointed /- @[simp] -theorem pointedToTwopSnd_comp_forget_to_bipointed : - pointedToTwopSnd ⋙ forget₂ Twop Bipointed = pointedToBipointedSnd := +theorem pointedToTwoPSnd_comp_forget_to_bipointed : + pointedToTwoPSnd ⋙ forget₂ TwoP Bipointed = pointedToBipointedSnd := rfl -#align Pointed_to_Twop_snd_comp_forget_to_Bipointed pointedToTwopSnd_comp_forget_to_bipointed +#align Pointed_to_Twop_snd_comp_forget_to_Bipointed pointedToTwoPSnd_comp_forget_to_bipointed +-/ +#print pointedToTwoPFstForgetCompBipointedToPointedFstAdjunction /- /-- Adding a second point is left adjoint to forgetting the second point. -/ -def pointedToTwopFstForgetCompBipointedToPointedFstAdjunction : - pointedToTwopFst ⊣ forget₂ Twop Bipointed ⋙ bipointedToPointedFst := +def pointedToTwoPFstForgetCompBipointedToPointedFstAdjunction : + pointedToTwoPFst ⊣ forget₂ TwoP Bipointed ⋙ bipointedToPointedFst := Adjunction.mkOfHomEquiv { homEquiv := fun X Y => { toFun := fun f => ⟨f.toFun ∘ Option.some, f.map_fst⟩ @@ -176,11 +221,13 @@ def pointedToTwopFstForgetCompBipointedToPointedFstAdjunction : by ext cases x <;> rfl } -#align Pointed_to_Twop_fst_forget_comp_Bipointed_to_Pointed_fst_adjunction pointedToTwopFstForgetCompBipointedToPointedFstAdjunction +#align Pointed_to_Twop_fst_forget_comp_Bipointed_to_Pointed_fst_adjunction pointedToTwoPFstForgetCompBipointedToPointedFstAdjunction +-/ +#print pointedToTwoPSndForgetCompBipointedToPointedSndAdjunction /- /-- Adding a first point is left adjoint to forgetting the first point. -/ -def pointedToTwopSndForgetCompBipointedToPointedSndAdjunction : - pointedToTwopSnd ⊣ forget₂ Twop Bipointed ⋙ bipointedToPointedSnd := +def pointedToTwoPSndForgetCompBipointedToPointedSndAdjunction : + pointedToTwoPSnd ⊣ forget₂ TwoP Bipointed ⋙ bipointedToPointedSnd := Adjunction.mkOfHomEquiv { homEquiv := fun X Y => { toFun := fun f => ⟨f.toFun ∘ Option.some, f.map_snd⟩ @@ -195,5 +242,6 @@ def pointedToTwopSndForgetCompBipointedToPointedSndAdjunction : by ext cases x <;> rfl } -#align Pointed_to_Twop_snd_forget_comp_Bipointed_to_Pointed_snd_adjunction pointedToTwopSndForgetCompBipointedToPointedSndAdjunction +#align Pointed_to_Twop_snd_forget_comp_Bipointed_to_Pointed_snd_adjunction pointedToTwoPSndForgetCompBipointedToPointedSndAdjunction +-/ diff --git a/Mathbin/CategoryTheory/DifferentialObject.lean b/Mathbin/CategoryTheory/DifferentialObject.lean index f970c87cba..a998465af9 100644 --- a/Mathbin/CategoryTheory/DifferentialObject.lean +++ b/Mathbin/CategoryTheory/DifferentialObject.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison ! This file was ported from Lean 3 source module category_theory.differential_object -! leanprover-community/mathlib commit f7707875544ef1f81b32cb68c79e0e24e45a0e76 +! leanprover-community/mathlib commit 6876fa15e3158ff3e4a4e2af1fb6e1945c6e8803 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ import Mathbin.Data.Int.Basic -import Mathbin.CategoryTheory.Shift +import Mathbin.CategoryTheory.Shift.Basic import Mathbin.CategoryTheory.ConcreteCategory.Basic /-! @@ -280,8 +280,9 @@ def shiftFunctor (n : ℤ) : DifferentialObject C ⥤ DifferentialObject C { f := f.f⟦n⟧' comm' := by dsimp - rw [category.assoc, shift_comm_hom_comp, ← functor.map_comp_assoc, f.comm, - functor.map_comp_assoc] } + erw [category.assoc, shift_comm_hom_comp, ← functor.map_comp_assoc, f.comm, + functor.map_comp_assoc] + rfl } map_id' := by intro X ext1 @@ -294,57 +295,58 @@ def shiftFunctor (n : ℤ) : DifferentialObject C ⥤ DifferentialObject C rw [functor.map_comp] #align category_theory.differential_object.shift_functor CategoryTheory.DifferentialObject.shiftFunctor -attribute [local simp] eq_to_hom_map - -attribute [local reducible] Discrete.addMonoidal shift_comm - /-- The shift functor on `differential_object C` is additive. -/ @[simps] def shiftFunctorAdd (m n : ℤ) : shiftFunctor C (m + n) ≅ shiftFunctor C m ⋙ shiftFunctor C n := by refine' nat_iso.of_components (fun X => mk_iso (shift_add X.pt _ _) _) _ · dsimp - -- This is just `simp, simp [eq_to_hom_map]`. - simp_rw [category.assoc, obj_μ_inv_app, μ_inv_hom_app_assoc, functor.map_comp, obj_μ_app, - category.assoc, μ_naturality_assoc, μ_inv_hom_app_assoc, obj_μ_inv_app, category.assoc, - μ_naturalityₗ_assoc, μ_inv_hom_app_assoc, μ_inv_naturalityᵣ_assoc] - simp only [eq_to_hom_map, eq_to_hom_app, eq_to_iso.hom, eq_to_hom_trans_assoc, eq_to_iso.inv] + rw [← cancel_epi ((shift_functor_add C m n).inv.app X.X)] + simp only [category.assoc, iso.inv_hom_id_app_assoc] + erw [← nat_trans.naturality_assoc] + dsimp + simp only [functor.map_comp, category.assoc, + shift_functor_comm_hom_app_comp_shift_shift_functor_add_hom_app 1 m n X.X, + iso.inv_hom_id_app_assoc] · intro X Y f ext dsimp exact nat_trans.naturality _ _ #align category_theory.differential_object.shift_functor_add CategoryTheory.DifferentialObject.shiftFunctorAdd -attribute [local reducible] endofunctor_monoidal_category - section -attribute [local instance] endofunctor_monoidal_category - /-- The shift by zero is naturally isomorphic to the identity. -/ @[simps] -def shiftε : 𝟭 (DifferentialObject C) ≅ shiftFunctor C 0 := +def shiftZero : shiftFunctor C 0 ≅ 𝟭 (DifferentialObject C) := by - refine' nat_iso.of_components (fun X => mk_iso ((shift_monoidal_functor C ℤ).εIso.app X.pt) _) _ - · dsimp - simp - dsimp - simp - · introv - ext + refine' nat_iso.of_components (fun X => mk_iso ((shift_functor_zero C ℤ).app X.pt) _) _ + · erw [← nat_trans.naturality] dsimp - simp -#align category_theory.differential_object.shift_ε CategoryTheory.DifferentialObject.shiftε + simp only [shift_functor_zero_hom_app_shift, category.assoc] + · tidy +#align category_theory.differential_object.shift_zero CategoryTheory.DifferentialObject.shiftZero end -attribute [local simp] eq_to_hom_map - instance : HasShift (DifferentialObject C) ℤ := hasShiftMk _ _ { f := shiftFunctor C - ε := shiftε C - μ := fun m n => (shiftFunctorAdd C m n).symm } + zero := shiftZero C + add := shiftFunctorAdd C + assoc_hom_app := fun m₁ m₂ m₃ X => by + ext1 + convert shift_functor_add_assoc_hom_app m₁ m₂ m₃ X.X + dsimp [shift_functor_add'] + simpa + zero_add_hom_app := fun n X => by + ext1 + convert shift_functor_add_zero_add_hom_app n X.X + simpa + add_zero_hom_app := fun n X => by + ext1 + convert shift_functor_add_add_zero_hom_app n X.X + simpa } end DifferentialObject diff --git a/Mathbin/CategoryTheory/GradedObject.lean b/Mathbin/CategoryTheory/GradedObject.lean index 2dd9115e62..b3b829bd85 100644 --- a/Mathbin/CategoryTheory/GradedObject.lean +++ b/Mathbin/CategoryTheory/GradedObject.lean @@ -4,13 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison ! This file was ported from Lean 3 source module category_theory.graded_object -! leanprover-community/mathlib commit 11613e2875d27371f380af8692498a22b66140a5 +! leanprover-community/mathlib commit 6876fa15e3158ff3e4a4e2af1fb6e1945c6e8803 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ import Mathbin.Algebra.GroupPower.Lemmas import Mathbin.CategoryTheory.Pi.Basic -import Mathbin.CategoryTheory.Shift +import Mathbin.CategoryTheory.Shift.Basic import Mathbin.CategoryTheory.ConcreteCategory.Basic /-! @@ -140,33 +140,30 @@ end instance hasShift {β : Type _} [AddCommGroup β] (s : β) : HasShift (GradedObjectWithShift s C) ℤ := hasShiftMk _ _ { f := fun n => comap (fun _ => C) fun b : β => b + n • s - ε := - (comapId β fun _ => C).symm ≪≫ - comapEq C + zero := + comapEq C (by ext - simp) - μ := fun m n => - comapComp _ _ _ ≪≫ - comapEq C + simp) ≪≫ + comapId β fun _ => C + add := fun m n => + comapEq C (by ext - simp [add_zsmul, add_comm]) - left_unitality := by - introv + simp [add_zsmul, add_comm]) ≪≫ + (comapComp _ _ _).symm + assoc_hom_app := fun m₁ m₂ m₃ X => by ext dsimp - simpa - right_unitality := by - introv + simp + zero_add_hom_app := fun n X => by ext dsimp simpa - associativity := by - introv + add_zero_hom_app := fun n X => by ext dsimp - simp } + simpa } #align category_theory.graded_object.has_shift CategoryTheory.GradedObject.hasShift @[simp] diff --git a/Mathbin/CategoryTheory/Idempotents/HomologicalComplex.lean b/Mathbin/CategoryTheory/Idempotents/HomologicalComplex.lean index 130bc32a1d..40c6f15faf 100644 --- a/Mathbin/CategoryTheory/Idempotents/HomologicalComplex.lean +++ b/Mathbin/CategoryTheory/Idempotents/HomologicalComplex.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou ! This file was ported from Lean 3 source module category_theory.idempotents.homological_complex -! leanprover-community/mathlib commit 31019c2504b17f85af7e0577585fad996935a317 +! leanprover-community/mathlib commit 200eda15d8ff5669854ff6bcc10aaf37cb70498f ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.Algebra.Homology.HomologicalComplex +import Mathbin.Algebra.Homology.Additive import Mathbin.CategoryTheory.Idempotents.Karoubi /-! @@ -18,6 +18,9 @@ This file contains simplifications lemmas for categories `karoubi (homological_complex C c)` and the construction of an equivalence of categories `karoubi (homological_complex C c) ≌ homological_complex (karoubi C) c`. +When the category `C` is idempotent complete, it is shown that +`homological_complex (karoubi C) c` is also idempotent complete. + -/ @@ -226,6 +229,13 @@ def karoubiCochainComplexEquivalence : karoubiHomologicalComplexEquivalence C (ComplexShape.up α) #align category_theory.idempotents.karoubi_cochain_complex_equivalence CategoryTheory.Idempotents.karoubiCochainComplexEquivalence +instance [IsIdempotentComplete C] : IsIdempotentComplete (HomologicalComplex C c) := + by + rw [is_idempotent_complete_iff_of_equivalence + ((to_karoubi_equivalence C).mapHomologicalComplex c), + ← is_idempotent_complete_iff_of_equivalence (karoubi_homological_complex_equivalence C c)] + infer_instance + end Idempotents end CategoryTheory diff --git a/Mathbin/CategoryTheory/Idempotents/Karoubi.lean b/Mathbin/CategoryTheory/Idempotents/Karoubi.lean index aee827ec80..be2d5052ba 100644 --- a/Mathbin/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathbin/CategoryTheory/Idempotents/Karoubi.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou ! This file was ported from Lean 3 source module category_theory.idempotents.karoubi -! leanprover-community/mathlib commit 88bca0ce5d22ebfd9e73e682e51d60ea13b48347 +! leanprover-community/mathlib commit 200eda15d8ff5669854ff6bcc10aaf37cb70498f ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -246,6 +246,17 @@ def toKaroubiIsEquivalence [IsIdempotentComplete C] : IsEquivalence (toKaroubi C Equivalence.ofFullyFaithfullyEssSurj (toKaroubi C) #align category_theory.idempotents.to_karoubi_is_equivalence CategoryTheory.Idempotents.toKaroubiIsEquivalence +/-- The equivalence `C ≅ karoubi C` when `C` is idempotent complete. -/ +def toKaroubiEquivalence [IsIdempotentComplete C] : C ≌ Karoubi C := + haveI := to_karoubi_is_equivalence C + functor.as_equivalence (to_karoubi C) +#align category_theory.idempotents.to_karoubi_equivalence CategoryTheory.Idempotents.toKaroubiEquivalence + +instance toKaroubiEquivalence_functor_additive [Preadditive C] [IsIdempotentComplete C] : + (toKaroubiEquivalence C).Functor.Additive := + (inferInstance : (toKaroubi C).Additive) +#align category_theory.idempotents.to_karoubi_equivalence_functor_additive CategoryTheory.Idempotents.toKaroubiEquivalence_functor_additive + namespace Karoubi variable {C} diff --git a/Mathbin/CategoryTheory/Shift.lean b/Mathbin/CategoryTheory/Shift.lean deleted file mode 100644 index d5d9c1e8c0..0000000000 --- a/Mathbin/CategoryTheory/Shift.lean +++ /dev/null @@ -1,557 +0,0 @@ -/- -Copyright (c) 2020 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison, Johan Commelin, Andrew Yang - -! This file was ported from Lean 3 source module category_theory.shift -! leanprover-community/mathlib commit 14b69e9f3c16630440a2cbd46f1ddad0d561dee7 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.CategoryTheory.Limits.Preserves.Shapes.Zero -import Mathbin.CategoryTheory.Monoidal.End -import Mathbin.CategoryTheory.Monoidal.Discrete - -/-! -# Shift - -A `shift` on a category `C` indexed by a monoid `A` is nothing more than a monoidal functor -from `A` to `C ⥤ C`. A typical example to keep in mind might be the category of -complexes `⋯ → C_{n-1} → C_n → C_{n+1} → ⋯`. It has a shift indexed by `ℤ`, where we assign to -each `n : ℤ` the functor `C ⥤ C` that re-indexes the terms, so the degree `i` term of `shift n C` -would be the degree `i+n`-th term of `C`. - -## Main definitions -* `has_shift`: A typeclass asserting the existence of a shift functor. -* `shift_equiv`: When the indexing monoid is a group, then the functor indexed by `n` and `-n` forms - an self-equivalence of `C`. -* `shift_comm`: When the indexing monoid is commutative, then shifts commute as well. - -## Implementation Notes - -Many of the definitions in this file are marked as an `abbreviation` so that the simp lemmas in -`category_theory/monoidal/End` can apply. - --/ - - -namespace CategoryTheory - -noncomputable section - -universe v u - -variable (C : Type u) (A : Type _) [Category.{v} C] - -attribute [local instance] endofunctor_monoidal_category - -section EqToHom - -variable {A C} - -variable [AddMonoid A] (F : MonoidalFunctor (Discrete A) (C ⥤ C)) - -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -@[simp, reassoc.1] -theorem eqToHom_μ_app {i j i' j' : A} (h₁ : i = i') (h₂ : j = j') (X : C) : - eqToHom (by rw [h₁, h₂] : (F.obj ⟨i⟩ ⊗ F.obj ⟨j⟩).obj X = (F.obj ⟨i'⟩ ⊗ F.obj ⟨j'⟩).obj X) ≫ - (F.μ ⟨i'⟩ ⟨j'⟩).app X = - (F.μ ⟨i⟩ ⟨j⟩).app X ≫ eqToHom (by rw [h₁, h₂]) := - by - cases h₁ - cases h₂ - rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] -#align category_theory.eq_to_hom_μ_app CategoryTheory.eqToHom_μ_app - -@[simp, reassoc.1] -theorem μ_inv_app_eqToHom {i j i' j' : A} (h₁ : i = i') (h₂ : j = j') (X : C) : - inv ((F.μ ⟨i⟩ ⟨j⟩).app X) ≫ eqToHom (by rw [h₁, h₂]) = - eqToHom (by rw [h₁, h₂]) ≫ inv ((F.μ ⟨i'⟩ ⟨j'⟩).app X) := - by - cases h₁ - cases h₂ - rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] -#align category_theory.μ_inv_app_eq_to_hom CategoryTheory.μ_inv_app_eqToHom - -end EqToHom - -variable {A C} - -/-- A monoidal functor from a group `A` into `C ⥤ C` induces -a self-equivalence of `C` for each `n : A`. -/ -@[simps Functor inverse unit_iso_hom unit_iso_inv counit_iso_hom counit_iso_inv] -def addNegEquiv [AddGroup A] (F : MonoidalFunctor (Discrete A) (C ⥤ C)) (n : A) : C ≌ C := - equivOfTensorIsoUnit F ⟨n⟩ ⟨(-n : A)⟩ (Discrete.eqToIso (add_neg_self n)) - (Discrete.eqToIso (neg_add_self n)) (Subsingleton.elim _ _) -#align category_theory.add_neg_equiv CategoryTheory.addNegEquiv - -section Defs - -variable (A C) [AddMonoid A] - -/-- A category has a shift indexed by an additive monoid `A` -if there is a monoidal functor from `A` to `C ⥤ C`. -/ -class HasShift (C : Type u) (A : Type _) [Category.{v} C] [AddMonoid A] where - shift : MonoidalFunctor (Discrete A) (C ⥤ C) -#align category_theory.has_shift CategoryTheory.HasShift - -/-- A helper structure to construct the shift functor `(discrete A) ⥤ (C ⥤ C)`. -/ -@[nolint has_nonempty_instance] -structure ShiftMkCore where - f : A → C ⥤ C - ε : 𝟭 C ≅ F 0 - μ : ∀ n m : A, F n ⋙ F m ≅ F (n + m) - associativity : - ∀ (m₁ m₂ m₃ : A) (X : C), - (F m₃).map ((μ m₁ m₂).Hom.app X) ≫ - (μ (m₁ + m₂) m₃).Hom.app X ≫ - eqToHom - (by - congr 2 - exact add_assoc _ _ _) = - (μ m₂ m₃).Hom.app ((F m₁).obj X) ≫ (μ m₁ (m₂ + m₃)).Hom.app X := by - obviously - left_unitality : - ∀ (n : A) (X : C), - (F n).map (ε.Hom.app X) ≫ (μ 0 n).Hom.app X = - eqToHom - (by - dsimp - rw [zero_add]) := by - obviously - right_unitality : - ∀ (n : A) (X : C), - ε.Hom.app ((F n).obj X) ≫ (μ n 0).Hom.app X = - eqToHom - (by - dsimp - rw [add_zero]) := by - obviously -#align category_theory.shift_mk_core CategoryTheory.ShiftMkCore - -section - -attribute [local simp] eq_to_hom_map - -attribute [local reducible] endofunctor_monoidal_category Discrete.addMonoidal - -/-- Constructs a `has_shift C A` instance from `shift_mk_core`. -/ -@[simps] -def hasShiftMk (h : ShiftMkCore C A) : HasShift C A := - ⟨{ Discrete.functor h.f with - ε := h.ε.Hom - μ := fun m n => (h.μ m.as n.as).Hom - μ_natural' := by - rintro ⟨X⟩ ⟨Y⟩ ⟨X'⟩ ⟨Y'⟩ ⟨⟨⟨rfl⟩⟩⟩ ⟨⟨⟨rfl⟩⟩⟩ - ext - dsimp - simp - dsimp - simp - associativity' := by - introv - ext - dsimp - simpa using h.associativity _ _ _ _ - left_unitality' := by - rintro ⟨X⟩ - ext - dsimp - rw [category.id_comp, ← category.assoc, h.left_unitality] - simp - right_unitality' := by - rintro ⟨X⟩ - ext - dsimp - rw [Functor.map_id, category.comp_id, ← category.assoc, h.right_unitality] - simp }⟩ -#align category_theory.has_shift_mk CategoryTheory.hasShiftMk - -end - -variable [HasShift C A] - -/-- The monoidal functor from `A` to `C ⥤ C` given a `has_shift` instance. -/ -def shiftMonoidalFunctor : MonoidalFunctor (Discrete A) (C ⥤ C) := - HasShift.shift -#align category_theory.shift_monoidal_functor CategoryTheory.shiftMonoidalFunctor - -variable {A} - -/-- The shift autoequivalence, moving objects and morphisms 'up'. -/ -abbrev shiftFunctor (i : A) : C ⥤ C := - (shiftMonoidalFunctor C A).obj ⟨i⟩ -#align category_theory.shift_functor CategoryTheory.shiftFunctor - -/-- Shifting by `i + j` is the same as shifting by `i` and then shifting by `j`. -/ -abbrev shiftFunctorAdd (i j : A) : shiftFunctor C (i + j) ≅ shiftFunctor C i ⋙ shiftFunctor C j := - ((shiftMonoidalFunctor C A).μIso ⟨i⟩ ⟨j⟩).symm -#align category_theory.shift_functor_add CategoryTheory.shiftFunctorAdd - -variable (A) - -/-- Shifting by zero is the identity functor. -/ -abbrev shiftFunctorZero : shiftFunctor C (0 : A) ≅ 𝟭 C := - (shiftMonoidalFunctor C A).εIso.symm -#align category_theory.shift_functor_zero CategoryTheory.shiftFunctorZero - --- mathport name: «expr ⟦ ⟧» -notation -- Any better notational suggestions? -X "⟦" n "⟧" => (shiftFunctor _ n).obj X - --- mathport name: «expr ⟦ ⟧'» -notation f "⟦" n "⟧'" => (shiftFunctor _ n).map f - -end Defs - -section AddMonoid - -variable {C A} [AddMonoid A] [HasShift C A] (X Y : C) (f : X ⟶ Y) - -@[simp] -theorem HasShift.shift_obj_obj (n : A) (X : C) : (HasShift.shift.obj ⟨n⟩).obj X = X⟦n⟧ := - rfl -#align category_theory.has_shift.shift_obj_obj CategoryTheory.HasShift.shift_obj_obj - -/-- Shifting by `i + j` is the same as shifting by `i` and then shifting by `j`. -/ -abbrev shiftAdd (i j : A) : X⟦i + j⟧ ≅ X⟦i⟧⟦j⟧ := - (shiftFunctorAdd C i j).app _ -#align category_theory.shift_add CategoryTheory.shiftAdd - -@[reassoc.1] -theorem shiftAdd_hom_comp_eq_to_hom₁ (i i' j : A) (h : i = i') : - (shiftAdd X i j).Hom ≫ eqToHom (by rw [h]) = eqToHom (by rw [h]) ≫ (shiftAdd X i' j).Hom := - by - cases h - rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] -#align category_theory.shift_add_hom_comp_eq_to_hom₁ CategoryTheory.shiftAdd_hom_comp_eq_to_hom₁ - -@[reassoc.1] -theorem shiftAdd_hom_comp_eq_to_hom₂ (i j j' : A) (h : j = j') : - (shiftAdd X i j).Hom ≫ eqToHom (by rw [h]) = eqToHom (by rw [h]) ≫ (shiftAdd X i j').Hom := - by - cases h - rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] -#align category_theory.shift_add_hom_comp_eq_to_hom₂ CategoryTheory.shiftAdd_hom_comp_eq_to_hom₂ - -@[reassoc.1] -theorem shiftAdd_hom_comp_eq_to_hom₁₂ (i j i' j' : A) (h₁ : i = i') (h₂ : j = j') : - (shiftAdd X i j).Hom ≫ eqToHom (by rw [h₁, h₂]) = - eqToHom (by rw [h₁, h₂]) ≫ (shiftAdd X i' j').Hom := - by - cases h₁ - cases h₂ - rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] -#align category_theory.shift_add_hom_comp_eq_to_hom₁₂ CategoryTheory.shiftAdd_hom_comp_eq_to_hom₁₂ - -@[reassoc.1] -theorem eqToHom_comp_shiftAdd_inv₁ (i i' j : A) (h : i = i') : - eqToHom (by rw [h]) ≫ (shiftAdd X i' j).inv = (shiftAdd X i j).inv ≫ eqToHom (by rw [h]) := by - rw [iso.comp_inv_eq, category.assoc, iso.eq_inv_comp, shift_add_hom_comp_eq_to_hom₁] -#align category_theory.eq_to_hom_comp_shift_add_inv₁ CategoryTheory.eqToHom_comp_shiftAdd_inv₁ - -@[reassoc.1] -theorem eqToHom_comp_shiftAdd_inv₂ (i j j' : A) (h : j = j') : - eqToHom (by rw [h]) ≫ (shiftAdd X i j').inv = (shiftAdd X i j).inv ≫ eqToHom (by rw [h]) := by - rw [iso.comp_inv_eq, category.assoc, iso.eq_inv_comp, shift_add_hom_comp_eq_to_hom₂] -#align category_theory.eq_to_hom_comp_shift_add_inv₂ CategoryTheory.eqToHom_comp_shiftAdd_inv₂ - -@[reassoc.1] -theorem eqToHom_comp_shiftAdd_inv₁₂ (i j i' j' : A) (h₁ : i = i') (h₂ : j = j') : - eqToHom (by rw [h₁, h₂]) ≫ (shiftAdd X i' j').inv = - (shiftAdd X i j).inv ≫ eqToHom (by rw [h₁, h₂]) := - by rw [iso.comp_inv_eq, category.assoc, iso.eq_inv_comp, shift_add_hom_comp_eq_to_hom₁₂] -#align category_theory.eq_to_hom_comp_shift_add_inv₁₂ CategoryTheory.eqToHom_comp_shiftAdd_inv₁₂ - -theorem shift_shift' (i j : A) : - f⟦i⟧'⟦j⟧' = (shiftAdd X i j).inv ≫ f⟦i + j⟧' ≫ (shiftAdd Y i j).Hom := - by - symm - apply nat_iso.naturality_1 -#align category_theory.shift_shift' CategoryTheory.shift_shift' - -variable (A) - -/-- Shifting by zero is the identity functor. -/ -abbrev shiftZero : X⟦0⟧ ≅ X := - (shiftFunctorZero C A).app _ -#align category_theory.shift_zero CategoryTheory.shiftZero - -theorem shift_zero' : f⟦(0 : A)⟧' = (shiftZero A X).Hom ≫ f ≫ (shiftZero A Y).inv := - by - symm - apply nat_iso.naturality_2 -#align category_theory.shift_zero' CategoryTheory.shift_zero' - -end AddMonoid - -section AddGroup - -variable (C) {A} [AddGroup A] [HasShift C A] - -variable (X Y : C) (f : X ⟶ Y) - -/-- Shifting by `i` is an equivalence. -/ -instance (i : A) : IsEquivalence (shiftFunctor C i) := - by - change is_equivalence (add_neg_equiv (shift_monoidal_functor C A) i).Functor - infer_instance - -@[simp] -theorem shiftFunctor_inv (i : A) : (shiftFunctor C i).inv = shiftFunctor C (-i) := - rfl -#align category_theory.shift_functor_inv CategoryTheory.shiftFunctor_inv - -/-- Shifting by `i` and then shifting by `-i` is the identity. -/ -abbrev shiftFunctorCompShiftFunctorNeg (i : A) : shiftFunctor C i ⋙ shiftFunctor C (-i) ≅ 𝟭 C := - unitOfTensorIsoUnit (shiftMonoidalFunctor C A) ⟨i⟩ ⟨(-i : A)⟩ (Discrete.eqToIso (add_neg_self i)) -#align category_theory.shift_functor_comp_shift_functor_neg CategoryTheory.shiftFunctorCompShiftFunctorNeg - -/-- Shifting by `-i` and then shifting by `i` is the identity. -/ -abbrev shiftFunctorNegCompShiftFunctor (i : A) : shiftFunctor C (-i) ⋙ shiftFunctor C i ≅ 𝟭 C := - unitOfTensorIsoUnit (shiftMonoidalFunctor C A) ⟨(-i : A)⟩ ⟨i⟩ (Discrete.eqToIso (neg_add_self i)) -#align category_theory.shift_functor_neg_comp_shift_functor CategoryTheory.shiftFunctorNegCompShiftFunctor - -section - -variable (C) - -/-- Shifting by `n` is a faithful functor. -/ -instance shiftFunctor_faithful (i : A) : Faithful (shiftFunctor C i) := - Faithful.of_comp_iso (shiftFunctorCompShiftFunctorNeg C i) -#align category_theory.shift_functor_faithful CategoryTheory.shiftFunctor_faithful - -/-- Shifting by `n` is a full functor. -/ -instance shiftFunctorFull (i : A) : Full (shiftFunctor C i) := - haveI : full (shift_functor C i ⋙ shift_functor C (-i)) := - full.of_iso (shift_functor_comp_shift_functor_neg C i).symm - full.of_comp_faithful _ (shift_functor C (-i)) -#align category_theory.shift_functor_full CategoryTheory.shiftFunctorFull - -/-- Shifting by `n` is an essentially surjective functor. -/ -instance shiftFunctor_essSurj (i : A) : EssSurj (shiftFunctor C i) - where mem_essImage Y := ⟨Y⟦-i⟧, ⟨(shiftFunctorNegCompShiftFunctor C i).app Y⟩⟩ -#align category_theory.shift_functor_ess_surj CategoryTheory.shiftFunctor_essSurj - -end - -variable {C} - -/-- Shifting by `i` and then shifting by `-i` is the identity. -/ -abbrev shiftShiftNeg (i : A) : X⟦i⟧⟦-i⟧ ≅ X := - (shiftFunctorCompShiftFunctorNeg C i).app _ -#align category_theory.shift_shift_neg CategoryTheory.shiftShiftNeg - -/-- Shifting by `-i` and then shifting by `i` is the identity. -/ -abbrev shiftNegShift (i : A) : X⟦-i⟧⟦i⟧ ≅ X := - (shiftFunctorNegCompShiftFunctor C i).app _ -#align category_theory.shift_neg_shift CategoryTheory.shiftNegShift - -variable {X Y} - -theorem shift_shift_neg' (i : A) : - f⟦i⟧'⟦-i⟧' = (shiftShiftNeg X i).Hom ≫ f ≫ (shiftShiftNeg Y i).inv := - by - symm - apply nat_iso.naturality_2 -#align category_theory.shift_shift_neg' CategoryTheory.shift_shift_neg' - -theorem shift_neg_shift' (i : A) : - f⟦-i⟧'⟦i⟧' = (shiftNegShift X i).Hom ≫ f ≫ (shiftNegShift Y i).inv := - by - symm - apply nat_iso.naturality_2 -#align category_theory.shift_neg_shift' CategoryTheory.shift_neg_shift' - -theorem shift_equiv_triangle (n : A) (X : C) : - (shiftShiftNeg X n).inv⟦n⟧' ≫ (shiftNegShift (X⟦n⟧) n).Hom = 𝟙 (X⟦n⟧) := - (addNegEquiv (shiftMonoidalFunctor C A) n).functor_unitIso_comp X -#align category_theory.shift_equiv_triangle CategoryTheory.shift_equiv_triangle - -section - -attribute [local reducible] Discrete.addMonoidal - -theorem shiftShiftNeg_hom_shift (n : A) (X : C) : - (shiftShiftNeg X n).Hom⟦n⟧' = (shiftNegShift (X⟦n⟧) n).Hom := - by - -- This is just `simp, simp [eq_to_hom_map]`. - simp only [iso.app_hom, unit_of_tensor_iso_unit_hom_app, eq_to_iso.hom, functor.map_comp, - obj_μ_app, eq_to_iso.inv, obj_ε_inv_app, μ_naturalityₗ_assoc, category.assoc, - μ_inv_hom_app_assoc, ε_inv_app_obj, μ_naturalityᵣ_assoc] - simp only [eq_to_hom_map, eq_to_hom_app, eq_to_hom_trans] -#align category_theory.shift_shift_neg_hom_shift CategoryTheory.shiftShiftNeg_hom_shift - -end - -theorem shiftShiftNeg_inv_shift (n : A) (X : C) : - (shiftShiftNeg X n).inv⟦n⟧' = (shiftNegShift (X⟦n⟧) n).inv := - by - ext - rw [← shift_shift_neg_hom_shift, ← functor.map_comp, iso.hom_inv_id, Functor.map_id] -#align category_theory.shift_shift_neg_inv_shift CategoryTheory.shiftShiftNeg_inv_shift - -@[simp] -theorem shiftShiftNeg_shift_eq (n : A) (X : C) : - (shiftFunctor C n).mapIso (shiftShiftNeg X n) = shiftNegShift (X⟦n⟧) n := - CategoryTheory.Iso.ext <| shiftShiftNeg_hom_shift _ _ -#align category_theory.shift_shift_neg_shift_eq CategoryTheory.shiftShiftNeg_shift_eq - -variable (C) - -/-- Shifting by `n` and shifting by `-n` forms an equivalence. -/ -@[simps] -def shiftEquiv (n : A) : C ≌ C := - { - addNegEquiv (shiftMonoidalFunctor C A) - n with - Functor := shiftFunctor C n - inverse := shiftFunctor C (-n) } -#align category_theory.shift_equiv CategoryTheory.shiftEquiv - -variable {C} - -open CategoryTheory.Limits - -variable [HasZeroMorphisms C] - -theorem shift_zero_eq_zero (X Y : C) (n : A) : (0 : X ⟶ Y)⟦n⟧' = (0 : X⟦n⟧ ⟶ Y⟦n⟧) := - CategoryTheory.Functor.map_zero _ _ _ -#align category_theory.shift_zero_eq_zero CategoryTheory.shift_zero_eq_zero - -end AddGroup - -section AddCommMonoid - -variable {C A} [AddCommMonoid A] [HasShift C A] - -variable (X Y : C) (f : X ⟶ Y) - -/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ -def shiftComm (i j : A) : X⟦i⟧⟦j⟧ ≅ X⟦j⟧⟦i⟧ := - (shiftAdd X i j).symm ≪≫ - ((shiftMonoidalFunctor C A).toFunctor.mapIso - (Discrete.eqToIso <| add_comm i j : (⟨i + j⟩ : Discrete A) ≅ ⟨j + i⟩)).app - X ≪≫ - shiftAdd X j i -#align category_theory.shift_comm CategoryTheory.shiftComm - -@[simp] -theorem shiftComm_symm (i j : A) : (shiftComm X i j).symm = shiftComm X j i := by ext; - dsimp [shift_comm]; simpa [eq_to_hom_map] -#align category_theory.shift_comm_symm CategoryTheory.shiftComm_symm - -variable {X Y} - -/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ -theorem shift_comm' (i j : A) : - f⟦i⟧'⟦j⟧' = (shiftComm _ _ _).Hom ≫ f⟦j⟧'⟦i⟧' ≫ (shiftComm _ _ _).Hom := - by - -- This is just `simp, simp [eq_to_hom_map]`. - simp only [shift_comm, iso.trans_hom, iso.symm_hom, iso.app_inv, iso.symm_inv, - monoidal_functor.μ_iso_hom, iso.app_hom, functor.map_iso_hom, eq_to_iso.hom, μ_naturality_assoc, - nat_trans.naturality_assoc, nat_trans.naturality, functor.comp_map, category.assoc, - μ_inv_hom_app_assoc] - simp only [eq_to_hom_map, eq_to_hom_app, eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp, - μ_hom_inv_app_assoc] -#align category_theory.shift_comm' CategoryTheory.shift_comm' - -@[reassoc.1] -theorem shiftComm_hom_comp (i j : A) : - (shiftComm X i j).Hom ≫ f⟦j⟧'⟦i⟧' = f⟦i⟧'⟦j⟧' ≫ (shiftComm Y i j).Hom := by - rw [shift_comm', ← shift_comm_symm, iso.symm_hom, iso.inv_hom_id_assoc] -#align category_theory.shift_comm_hom_comp CategoryTheory.shiftComm_hom_comp - -end AddCommMonoid - -variable {D : Type _} [Category D] [AddMonoid A] [HasShift D A] - -variable (F : C ⥤ D) [Full F] [Faithful F] - -section - -attribute [local reducible] Discrete.addMonoidal - -/-- Given a family of endomorphisms of `C` which are interwined by a fully faithful `F : C ⥤ D` -with shift functors on `D`, we can promote that family to shift functors on `C`. -/ -def hasShiftOfFullyFaithful (s : A → C ⥤ C) (i : ∀ i, s i ⋙ F ≅ F ⋙ shiftFunctor D i) : - HasShift C A := - hasShiftMk C A - { f := s - ε := - natIsoOfCompFullyFaithful F - (calc - 𝟭 C ⋙ F ≅ F := Functor.leftUnitor _ - _ ≅ F ⋙ 𝟭 D := (Functor.rightUnitor _).symm - _ ≅ F ⋙ shiftFunctor D (0 : A) := (isoWhiskerLeft F (shiftFunctorZero D A).symm) - _ ≅ s 0 ⋙ F := (i 0).symm - ) - μ := fun a b => - natIsoOfCompFullyFaithful F - (calc - (s a ⋙ s b) ⋙ F ≅ s a ⋙ s b ⋙ F := Functor.associator _ _ _ - _ ≅ s a ⋙ F ⋙ shiftFunctor D b := (isoWhiskerLeft _ (i b)) - _ ≅ (s a ⋙ F) ⋙ shiftFunctor D b := (Functor.associator _ _ _).symm - _ ≅ (F ⋙ shiftFunctor D a) ⋙ shiftFunctor D b := (isoWhiskerRight (i a) _) - _ ≅ F ⋙ shiftFunctor D a ⋙ shiftFunctor D b := (Functor.associator _ _ _) - _ ≅ F ⋙ shiftFunctor D (a + b) := (isoWhiskerLeft _ (shiftFunctorAdd D a b).symm) - _ ≅ s (a + b) ⋙ F := (i (a + b)).symm - ) - associativity := by - intros ; apply F.map_injective; dsimp - simp only [category.comp_id, category.id_comp, category.assoc, - CategoryTheory.Functor.map_comp, functor.image_preimage, eq_to_hom_map, - iso.inv_hom_id_app_assoc] - erw [(i m₃).Hom.naturality_assoc] - congr 1 - dsimp - simp only [eq_to_iso.inv, eq_to_hom_app, eq_to_hom_map, obj_μ_app, μ_naturality_assoc, - category.assoc, CategoryTheory.Functor.map_comp, functor.image_preimage] - congr 3 - dsimp - simp only [← (shift_functor D m₃).map_comp_assoc, iso.inv_hom_id_app] - erw [(shift_functor D m₃).map_id, category.id_comp] - erw [((shift_monoidal_functor D A).μIso ⟨m₁ + m₂⟩ ⟨m₃⟩).inv_hom_id_app_assoc] - congr 1 - have := dcongr_arg (fun a => (i a).inv.app X) (add_assoc m₁ m₂ m₃) - dsimp at this - simp [this] - left_unitality := by - intros ; apply F.map_injective; dsimp - simp only [category.comp_id, category.id_comp, category.assoc, - CategoryTheory.Functor.map_comp, eq_to_hom_app, eq_to_hom_map, functor.image_preimage] - erw [(i n).Hom.naturality_assoc] - dsimp - simp only [eq_to_iso.inv, eq_to_hom_app, category.assoc, CategoryTheory.Functor.map_comp, - eq_to_hom_map, obj_ε_app, functor.image_preimage] - simp only [← (shift_functor D n).map_comp_assoc, iso.inv_hom_id_app] - dsimp - simp only [category.id_comp, μ_inv_hom_app_assoc, CategoryTheory.Functor.map_id] - have := dcongr_arg (fun a => (i a).inv.app X) (zero_add n) - dsimp at this - simp [this] - right_unitality := by - intros ; apply F.map_injective; dsimp - simp only [category.comp_id, category.id_comp, category.assoc, iso.inv_hom_id_app_assoc, - eq_to_iso.inv, eq_to_hom_app, eq_to_hom_map, CategoryTheory.Functor.map_comp, - functor.image_preimage, obj_zero_map_μ_app, ε_hom_inv_app_assoc] - have := dcongr_arg (fun a => (i a).inv.app X) (add_zero n) - dsimp at this - simp [this] } -#align category_theory.has_shift_of_fully_faithful CategoryTheory.hasShiftOfFullyFaithful - -end - --- incorrectly reports that `[full F]` and `[faithful F]` are unused. -/-- When we construct shifts on a subcategory from shifts on the ambient category, -the inclusion functor intertwines the shifts. -/ -@[nolint unused_arguments] -def hasShiftOfFullyFaithfulComm (s : A → C ⥤ C) (i : ∀ i, s i ⋙ F ≅ F ⋙ shiftFunctor D i) (m : A) : - haveI := has_shift_of_fully_faithful F s i - shift_functor C m ⋙ F ≅ F ⋙ shift_functor D m := - i m -#align category_theory.has_shift_of_fully_faithful_comm CategoryTheory.hasShiftOfFullyFaithfulComm - -end CategoryTheory - diff --git a/Mathbin/CategoryTheory/Shift/Basic.lean b/Mathbin/CategoryTheory/Shift/Basic.lean new file mode 100644 index 0000000000..3ba85270b3 --- /dev/null +++ b/Mathbin/CategoryTheory/Shift/Basic.lean @@ -0,0 +1,818 @@ +/- +Copyright (c) 2020 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison, Johan Commelin, Andrew Yang, Joël Riou + +! This file was ported from Lean 3 source module category_theory.shift.basic +! leanprover-community/mathlib commit 6876fa15e3158ff3e4a4e2af1fb6e1945c6e8803 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Limits.Preserves.Shapes.Zero +import Mathbin.CategoryTheory.Monoidal.End +import Mathbin.CategoryTheory.Monoidal.Discrete + +/-! +# Shift + +A `shift` on a category `C` indexed by a monoid `A` is nothing more than a monoidal functor +from `A` to `C ⥤ C`. A typical example to keep in mind might be the category of +complexes `⋯ → C_{n-1} → C_n → C_{n+1} → ⋯`. It has a shift indexed by `ℤ`, where we assign to +each `n : ℤ` the functor `C ⥤ C` that re-indexes the terms, so the degree `i` term of `shift n C` +would be the degree `i+n`-th term of `C`. + +## Main definitions +* `has_shift`: A typeclass asserting the existence of a shift functor. +* `shift_equiv`: When the indexing monoid is a group, then the functor indexed by `n` and `-n` forms + an self-equivalence of `C`. +* `shift_comm`: When the indexing monoid is commutative, then shifts commute as well. + +## Implementation Notes + +`[has_shift C A]` is implemented using `monoidal_functor (discrete A) (C ⥤ C)`. +However, the API of monodial functors is used only internally: one should use the API of +shifts functors which includes `shift_functor C a : C ⥤ C` for `a : A`, +`shift_functor_zero C A : shift_functor C (0 : A) ≅ 𝟭 C` and +`shift_functor_add C i j : shift_functor C (i + j) ≅ shift_functor C i ⋙ shift_functor C j` +(and its variant `shift_functor_add'`). These isomorphisms satisfy some coherence properties +which are stated in lemmas like `shift_functor_add'_assoc`, `shift_functor_add'_zero_add` and +`shift_functor_add'_add_zero`. + +-/ + + +namespace CategoryTheory + +noncomputable section + +universe v u + +variable (C : Type u) (A : Type _) [Category.{v} C] + +attribute [local instance] endofunctor_monoidal_category + +section Defs + +variable (A C) [AddMonoid A] + +/-- A category has a shift indexed by an additive monoid `A` +if there is a monoidal functor from `A` to `C ⥤ C`. -/ +class HasShift (C : Type u) (A : Type _) [Category.{v} C] [AddMonoid A] where + shift : MonoidalFunctor (Discrete A) (C ⥤ C) +#align category_theory.has_shift CategoryTheory.HasShift + +/-- A helper structure to construct the shift functor `(discrete A) ⥤ (C ⥤ C)`. -/ +@[nolint has_nonempty_instance] +structure ShiftMkCore where + f : A → C ⥤ C + zero : F 0 ≅ 𝟭 C + add : ∀ n m : A, F (n + m) ≅ F n ⋙ F m + assoc_hom_app : + ∀ (m₁ m₂ m₃ : A) (X : C), + (add (m₁ + m₂) m₃).Hom.app X ≫ (F m₃).map ((add m₁ m₂).Hom.app X) = + eqToHom (by rw [add_assoc]) ≫ + (add m₁ (m₂ + m₃)).Hom.app X ≫ (add m₂ m₃).Hom.app ((F m₁).obj X) + zero_add_hom_app : + ∀ (n : A) (X : C), + (add 0 n).Hom.app X = eqToHom (by dsimp <;> rw [zero_add]) ≫ (F n).map (zero.inv.app X) + add_zero_hom_app : + ∀ (n : A) (X : C), + (add n 0).Hom.app X = eqToHom (by dsimp <;> rw [add_zero]) ≫ zero.inv.app ((F n).obj X) +#align category_theory.shift_mk_core CategoryTheory.ShiftMkCore + +namespace ShiftMkCore + +variable {C A} + +attribute [reassoc.1] assoc_hom_app + +@[reassoc.1] +theorem assoc_inv_app (h : ShiftMkCore C A) (m₁ m₂ m₃ : A) (X : C) : + (h.f m₃).map ((h.add m₁ m₂).inv.app X) ≫ (h.add (m₁ + m₂) m₃).inv.app X = + (h.add m₂ m₃).inv.app ((h.f m₁).obj X) ≫ + (h.add m₁ (m₂ + m₃)).inv.app X ≫ eqToHom (by rw [add_assoc]) := + by + rw [← cancel_mono ((h.add (m₁ + m₂) m₃).Hom.app X ≫ (h.F m₃).map ((h.add m₁ m₂).Hom.app X)), + category.assoc, category.assoc, category.assoc, iso.inv_hom_id_app_assoc, ← functor.map_comp, + iso.inv_hom_id_app, Functor.map_id, h.assoc_hom_app, eq_to_hom_trans_assoc, eq_to_hom_refl, + category.id_comp, iso.inv_hom_id_app_assoc, iso.inv_hom_id_app] + rfl +#align category_theory.shift_mk_core.assoc_inv_app CategoryTheory.ShiftMkCore.assoc_inv_app + +theorem zero_add_inv_app (h : ShiftMkCore C A) (n : A) (X : C) : + (h.add 0 n).inv.app X = (h.f n).map (h.zero.Hom.app X) ≫ eqToHom (by dsimp <;> rw [zero_add]) := + by + rw [← cancel_epi ((h.add 0 n).Hom.app X), iso.hom_inv_id_app, h.zero_add_hom_app, category.assoc, + ← functor.map_comp_assoc, iso.inv_hom_id_app, Functor.map_id, category.id_comp, eq_to_hom_trans, + eq_to_hom_refl] +#align category_theory.shift_mk_core.zero_add_inv_app CategoryTheory.ShiftMkCore.zero_add_inv_app + +theorem add_zero_inv_app (h : ShiftMkCore C A) (n : A) (X : C) : + (h.add n 0).inv.app X = h.zero.Hom.app ((h.f n).obj X) ≫ eqToHom (by dsimp <;> rw [add_zero]) := + by + rw [← cancel_epi ((h.add n 0).Hom.app X), iso.hom_inv_id_app, h.add_zero_hom_app, category.assoc, + iso.inv_hom_id_app_assoc, eq_to_hom_trans, eq_to_hom_refl] +#align category_theory.shift_mk_core.add_zero_inv_app CategoryTheory.ShiftMkCore.add_zero_inv_app + +end ShiftMkCore + +section + +attribute [local simp] eq_to_hom_map + +attribute [local reducible] endofunctor_monoidal_category Discrete.addMonoidal + +/-- Constructs a `has_shift C A` instance from `shift_mk_core`. -/ +def hasShiftMk (h : ShiftMkCore C A) : HasShift C A := + ⟨{ Discrete.functor h.f with + ε := h.zero.inv + μ := fun m n => (h.add m.as n.as).inv + μ_natural' := by + rintro ⟨X⟩ ⟨Y⟩ ⟨X'⟩ ⟨Y'⟩ ⟨⟨⟨rfl⟩⟩⟩ ⟨⟨⟨rfl⟩⟩⟩ + ext + dsimp + simp only [discrete.functor_map_id, category.assoc] + dsimp + simp + associativity' := by + rintro ⟨m₁⟩ ⟨m₂⟩ ⟨m₃⟩ + ext X + dsimp + simp [h.assoc_inv_app_assoc m₁ m₂ m₃ X] + left_unitality' := by + rintro ⟨n⟩ + ext X + dsimp + simp only [h.zero_add_inv_app, ← functor.map_comp, category.id_comp, eq_to_hom_map, + eq_to_hom_app, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id, + iso.inv_hom_id_app] + erw [Functor.map_id] + right_unitality' := by + rintro ⟨n⟩ + ext X + dsimp + simpa only [h.add_zero_inv_app, Functor.map_id, category.comp_id, eq_to_hom_map, + eq_to_hom_app, category.assoc, eq_to_hom_trans, eq_to_hom_refl, iso.inv_hom_id_app] }⟩ +#align category_theory.has_shift_mk CategoryTheory.hasShiftMk + +end + +section + +variable [HasShift C A] + +/-- The monoidal functor from `A` to `C ⥤ C` given a `has_shift` instance. -/ +def shiftMonoidalFunctor : MonoidalFunctor (Discrete A) (C ⥤ C) := + HasShift.shift +#align category_theory.shift_monoidal_functor CategoryTheory.shiftMonoidalFunctor + +variable {A} + +/-- The shift autoequivalence, moving objects and morphisms 'up'. -/ +def shiftFunctor (i : A) : C ⥤ C := + (shiftMonoidalFunctor C A).obj ⟨i⟩ +#align category_theory.shift_functor CategoryTheory.shiftFunctor + +/-- Shifting by `i + j` is the same as shifting by `i` and then shifting by `j`. -/ +def shiftFunctorAdd (i j : A) : shiftFunctor C (i + j) ≅ shiftFunctor C i ⋙ shiftFunctor C j := + ((shiftMonoidalFunctor C A).μIso ⟨i⟩ ⟨j⟩).symm +#align category_theory.shift_functor_add CategoryTheory.shiftFunctorAdd + +/-- When `k = i + j`, shifting by `k` is the same as shifting by `i` and then shifting by `j`. -/ +def shiftFunctorAdd' (i j k : A) (h : i + j = k) : + shiftFunctor C k ≅ shiftFunctor C i ⋙ shiftFunctor C j := + eqToIso (by rw [h]) ≪≫ shiftFunctorAdd C i j +#align category_theory.shift_functor_add' CategoryTheory.shiftFunctorAdd' + +theorem shiftFunctorAdd'_eq_shiftFunctorAdd (i j : A) : + shiftFunctorAdd' C i j (i + j) rfl = shiftFunctorAdd C i j := + by + ext1 + apply category.id_comp +#align category_theory.shift_functor_add'_eq_shift_functor_add CategoryTheory.shiftFunctorAdd'_eq_shiftFunctorAdd + +variable (A) + +/-- Shifting by zero is the identity functor. -/ +def shiftFunctorZero : shiftFunctor C (0 : A) ≅ 𝟭 C := + (shiftMonoidalFunctor C A).εIso.symm +#align category_theory.shift_functor_zero CategoryTheory.shiftFunctorZero + +end + +variable {C A} + +theorem ShiftMkCore.shiftFunctor_eq (h : ShiftMkCore C A) (a : A) : + @shiftFunctor _ _ _ _ (hasShiftMk C A h) a = h.f a := + Functor.ext (by tidy) (by tidy) +#align category_theory.shift_mk_core.shift_functor_eq CategoryTheory.ShiftMkCore.shiftFunctor_eq + +theorem ShiftMkCore.shiftFunctorZero_eq (h : ShiftMkCore C A) : + @shiftFunctorZero _ _ _ _ (hasShiftMk C A h) = h.zero := + by + letI := has_shift_mk C A h + ext1 + suffices (shift_functor_zero C A).inv = h.zero.inv + by + rw [← cancel_mono h.zero.inv, h.zero.hom_inv_id, ← this, iso.hom_inv_id] + rfl + rfl +#align category_theory.shift_mk_core.shift_functor_zero_eq CategoryTheory.ShiftMkCore.shiftFunctorZero_eq + +theorem ShiftMkCore.shiftFunctorAdd_eq (h : ShiftMkCore C A) (a b : A) : + @shiftFunctorAdd _ _ _ _ (hasShiftMk C A h) a b = h.add a b := + by + letI := has_shift_mk C A h + ext1 + suffices (shift_functor_add C a b).inv = (h.add a b).inv + by + rw [← cancel_mono (h.add a b).inv, (h.add a b).hom_inv_id, ← this, iso.hom_inv_id] + rfl + rfl +#align category_theory.shift_mk_core.shift_functor_add_eq CategoryTheory.ShiftMkCore.shiftFunctorAdd_eq + +-- mathport name: «expr ⟦ ⟧» +notation -- Any better notational suggestions? +X "⟦" n "⟧" => (shiftFunctor _ n).obj X + +-- mathport name: «expr ⟦ ⟧'» +notation f "⟦" n "⟧'" => (shiftFunctor _ n).map f + +variable (C) + +variable [HasShift C A] + +attribute [local reducible] endofunctor_monoidal_category Discrete.addMonoidal + +theorem shiftFunctorAdd'_zero_add (a : A) : + shiftFunctorAdd' C 0 a a (zero_add a) = + (Functor.leftUnitor _).symm ≪≫ + isoWhiskerRight (shiftFunctorZero C A).symm (shiftFunctor C a) := + by + ext X + dsimp [shift_functor_add'] + erw [obj_ε_app] + simpa [eq_to_hom_map] +#align category_theory.shift_functor_add'_zero_add CategoryTheory.shiftFunctorAdd'_zero_add + +theorem shiftFunctorAdd'_add_zero (a : A) : + shiftFunctorAdd' C a 0 a (add_zero a) = + (Functor.rightUnitor _).symm ≪≫ + isoWhiskerLeft (shiftFunctor C a) (shiftFunctorZero C A).symm := + by + ext X + dsimp [shift_functor_add'] + erw [ε_app_obj] + simpa [eq_to_hom_map] +#align category_theory.shift_functor_add'_add_zero CategoryTheory.shiftFunctorAdd'_add_zero + +theorem shiftFunctorAdd'_assoc (a₁ a₂ a₃ a₁₂ a₂₃ a₁₂₃ : A) (h₁₂ : a₁ + a₂ = a₁₂) + (h₂₃ : a₂ + a₃ = a₂₃) (h₁₂₃ : a₁ + a₂ + a₃ = a₁₂₃) : + shiftFunctorAdd' C a₁₂ a₃ a₁₂₃ (by rw [← h₁₂, h₁₂₃]) ≪≫ + isoWhiskerRight (shiftFunctorAdd' C a₁ a₂ a₁₂ h₁₂) _ ≪≫ Functor.associator _ _ _ = + shiftFunctorAdd' C a₁ a₂₃ a₁₂₃ (by rw [← h₂₃, ← add_assoc, h₁₂₃]) ≪≫ + isoWhiskerLeft _ (shiftFunctorAdd' C a₂ a₃ a₂₃ h₂₃) := + by + substs h₁₂ h₂₃ h₁₂₃ + ext X + dsimp + simp only [shift_functor_add'_eq_shift_functor_add, category.comp_id] + dsimp [shift_functor_add', shift_functor_add, shift_functor] + simp [obj_μ_inv_app, eq_to_hom_map] +#align category_theory.shift_functor_add'_assoc CategoryTheory.shiftFunctorAdd'_assoc + +theorem shiftFunctorAdd_assoc (a₁ a₂ a₃ : A) : + shiftFunctorAdd C (a₁ + a₂) a₃ ≪≫ + isoWhiskerRight (shiftFunctorAdd C a₁ a₂) _ ≪≫ Functor.associator _ _ _ = + shiftFunctorAdd' C a₁ (a₂ + a₃) _ (add_assoc a₁ a₂ a₃).symm ≪≫ + isoWhiskerLeft _ (shiftFunctorAdd C a₂ a₃) := + by + ext X + simpa [shift_functor_add'_eq_shift_functor_add] using + nat_trans.congr_app (congr_arg iso.hom (shift_functor_add'_assoc C a₁ a₂ a₃ _ _ _ rfl rfl rfl)) + X +#align category_theory.shift_functor_add_assoc CategoryTheory.shiftFunctorAdd_assoc + +variable {C} + +theorem shiftFunctorAdd'_zero_add_hom_app (a : A) (X : C) : + (shiftFunctorAdd' C 0 a a (zero_add a)).Hom.app X = (shiftFunctorZero C A).inv.app X⟦a⟧' := by + simpa using nat_trans.congr_app (congr_arg iso.hom (shift_functor_add'_zero_add C a)) X +#align category_theory.shift_functor_add'_zero_add_hom_app CategoryTheory.shiftFunctorAdd'_zero_add_hom_app + +theorem shiftFunctorAdd_zero_add_hom_app (a : A) (X : C) : + (shiftFunctorAdd C 0 a).Hom.app X = + eqToHom (by dsimp <;> rw [zero_add]) ≫ (shiftFunctorZero C A).inv.app X⟦a⟧' := + by + erw [← shift_functor_add'_zero_add_hom_app] + dsimp [shift_functor_add'] + simp +#align category_theory.shift_functor_add_zero_add_hom_app CategoryTheory.shiftFunctorAdd_zero_add_hom_app + +theorem shiftFunctorAdd'_zero_add_inv_app (a : A) (X : C) : + (shiftFunctorAdd' C 0 a a (zero_add a)).inv.app X = (shiftFunctorZero C A).Hom.app X⟦a⟧' := + by + have := nat_trans.congr_app (congr_arg iso.inv (shift_functor_add'_zero_add C a)) X + simp only [iso.trans_inv, iso_whisker_right_inv, iso.symm_inv, nat_trans.comp_app, + whisker_right_app, functor.left_unitor_hom_app] at this + dsimp at this + simpa only [category.comp_id] using this +#align category_theory.shift_functor_add'_zero_add_inv_app CategoryTheory.shiftFunctorAdd'_zero_add_inv_app + +theorem shiftFunctorAdd_zero_add_inv_app (a : A) (X : C) : + (shiftFunctorAdd C 0 a).inv.app X = + (shiftFunctorZero C A).Hom.app X⟦a⟧' ≫ eqToHom (by dsimp <;> rw [zero_add]) := + by + erw [← shift_functor_add'_zero_add_inv_app] + dsimp [shift_functor_add'] + simp +#align category_theory.shift_functor_add_zero_add_inv_app CategoryTheory.shiftFunctorAdd_zero_add_inv_app + +theorem shiftFunctorAdd'_add_zero_hom_app (a : A) (X : C) : + (shiftFunctorAdd' C a 0 a (add_zero a)).Hom.app X = (shiftFunctorZero C A).inv.app (X⟦a⟧) := by + simpa using nat_trans.congr_app (congr_arg iso.hom (shift_functor_add'_add_zero C a)) X +#align category_theory.shift_functor_add'_add_zero_hom_app CategoryTheory.shiftFunctorAdd'_add_zero_hom_app + +theorem shiftFunctorAdd_add_zero_hom_app (a : A) (X : C) : + (shiftFunctorAdd C a 0).Hom.app X = + eqToHom (by dsimp <;> rw [add_zero]) ≫ (shiftFunctorZero C A).inv.app (X⟦a⟧) := + by + rw [← shift_functor_add'_add_zero_hom_app] + dsimp [shift_functor_add'] + simp +#align category_theory.shift_functor_add_add_zero_hom_app CategoryTheory.shiftFunctorAdd_add_zero_hom_app + +theorem shiftFunctorAdd'_add_zero_inv_app (a : A) (X : C) : + (shiftFunctorAdd' C a 0 a (add_zero a)).inv.app X = (shiftFunctorZero C A).Hom.app (X⟦a⟧) := + by + have := nat_trans.congr_app (congr_arg iso.inv (shift_functor_add'_add_zero C a)) X + simp only [iso.trans_inv, iso_whisker_left_inv, iso.symm_inv, nat_trans.comp_app, + whisker_left_app, functor.right_unitor_hom_app] at this + dsimp at this + simpa only [category.comp_id] using this +#align category_theory.shift_functor_add'_add_zero_inv_app CategoryTheory.shiftFunctorAdd'_add_zero_inv_app + +theorem shiftFunctorAdd_add_zero_inv_app (a : A) (X : C) : + (shiftFunctorAdd C a 0).inv.app X = + (shiftFunctorZero C A).Hom.app (X⟦a⟧) ≫ eqToHom (by dsimp <;> rw [add_zero]) := + by + rw [← shift_functor_add'_add_zero_inv_app] + dsimp [shift_functor_add'] + simp +#align category_theory.shift_functor_add_add_zero_inv_app CategoryTheory.shiftFunctorAdd_add_zero_inv_app + +@[reassoc.1] +theorem shiftFunctorAdd'_assoc_hom_app (a₁ a₂ a₃ a₁₂ a₂₃ a₁₂₃ : A) (h₁₂ : a₁ + a₂ = a₁₂) + (h₂₃ : a₂ + a₃ = a₂₃) (h₁₂₃ : a₁ + a₂ + a₃ = a₁₂₃) (X : C) : + (shiftFunctorAdd' C a₁₂ a₃ a₁₂₃ (by rw [← h₁₂, h₁₂₃])).Hom.app X ≫ + (shiftFunctorAdd' C a₁ a₂ a₁₂ h₁₂).Hom.app X⟦a₃⟧' = + (shiftFunctorAdd' C a₁ a₂₃ a₁₂₃ (by rw [← h₂₃, ← add_assoc, h₁₂₃])).Hom.app X ≫ + (shiftFunctorAdd' C a₂ a₃ a₂₃ h₂₃).Hom.app (X⟦a₁⟧) := + by + simpa using + nat_trans.congr_app (congr_arg iso.hom (shift_functor_add'_assoc C _ _ _ _ _ _ h₁₂ h₂₃ h₁₂₃)) X +#align category_theory.shift_functor_add'_assoc_hom_app CategoryTheory.shiftFunctorAdd'_assoc_hom_app + +@[reassoc.1] +theorem shiftFunctorAdd'_assoc_inv_app (a₁ a₂ a₃ a₁₂ a₂₃ a₁₂₃ : A) (h₁₂ : a₁ + a₂ = a₁₂) + (h₂₃ : a₂ + a₃ = a₂₃) (h₁₂₃ : a₁ + a₂ + a₃ = a₁₂₃) (X : C) : + (shiftFunctorAdd' C a₁ a₂ a₁₂ h₁₂).inv.app X⟦a₃⟧' ≫ + (shiftFunctorAdd' C a₁₂ a₃ a₁₂₃ (by rw [← h₁₂, h₁₂₃])).inv.app X = + (shiftFunctorAdd' C a₂ a₃ a₂₃ h₂₃).inv.app (X⟦a₁⟧) ≫ + (shiftFunctorAdd' C a₁ a₂₃ a₁₂₃ (by rw [← h₂₃, ← add_assoc, h₁₂₃])).inv.app X := + by + simpa using + nat_trans.congr_app (congr_arg iso.inv (shift_functor_add'_assoc C _ _ _ _ _ _ h₁₂ h₂₃ h₁₂₃)) X +#align category_theory.shift_functor_add'_assoc_inv_app CategoryTheory.shiftFunctorAdd'_assoc_inv_app + +@[reassoc.1] +theorem shiftFunctorAdd_assoc_hom_app (a₁ a₂ a₃ : A) (X : C) : + (shiftFunctorAdd C (a₁ + a₂) a₃).Hom.app X ≫ (shiftFunctorAdd C a₁ a₂).Hom.app X⟦a₃⟧' = + (shiftFunctorAdd' C a₁ (a₂ + a₃) (a₁ + a₂ + a₃) (add_assoc _ _ _).symm).Hom.app X ≫ + (shiftFunctorAdd C a₂ a₃).Hom.app (X⟦a₁⟧) := + by simpa using nat_trans.congr_app (congr_arg iso.hom (shift_functor_add_assoc C a₁ a₂ a₃)) X +#align category_theory.shift_functor_add_assoc_hom_app CategoryTheory.shiftFunctorAdd_assoc_hom_app + +@[reassoc.1] +theorem shiftFunctorAdd_assoc_inv_app (a₁ a₂ a₃ : A) (X : C) : + (shiftFunctorAdd C a₁ a₂).inv.app X⟦a₃⟧' ≫ (shiftFunctorAdd C (a₁ + a₂) a₃).inv.app X = + (shiftFunctorAdd C a₂ a₃).inv.app (X⟦a₁⟧) ≫ + (shiftFunctorAdd' C a₁ (a₂ + a₃) (a₁ + a₂ + a₃) (add_assoc _ _ _).symm).inv.app X := + by simpa using nat_trans.congr_app (congr_arg iso.inv (shift_functor_add_assoc C a₁ a₂ a₃)) X +#align category_theory.shift_functor_add_assoc_inv_app CategoryTheory.shiftFunctorAdd_assoc_inv_app + +end Defs + +section AddMonoid + +variable {C A} [AddMonoid A] [HasShift C A] (X Y : C) (f : X ⟶ Y) + +@[simp] +theorem HasShift.shift_obj_obj (n : A) (X : C) : (HasShift.shift.obj ⟨n⟩).obj X = X⟦n⟧ := + rfl +#align category_theory.has_shift.shift_obj_obj CategoryTheory.HasShift.shift_obj_obj + +/-- Shifting by `i + j` is the same as shifting by `i` and then shifting by `j`. -/ +abbrev shiftAdd (i j : A) : X⟦i + j⟧ ≅ X⟦i⟧⟦j⟧ := + (shiftFunctorAdd C i j).app _ +#align category_theory.shift_add CategoryTheory.shiftAdd + +theorem shift_shift' (i j : A) : + f⟦i⟧'⟦j⟧' = (shiftAdd X i j).inv ≫ f⟦i + j⟧' ≫ (shiftAdd Y i j).Hom := + by + symm + apply nat_iso.naturality_1 +#align category_theory.shift_shift' CategoryTheory.shift_shift' + +variable (A) + +/-- Shifting by zero is the identity functor. -/ +abbrev shiftZero : X⟦0⟧ ≅ X := + (shiftFunctorZero C A).app _ +#align category_theory.shift_zero CategoryTheory.shiftZero + +theorem shift_zero' : f⟦(0 : A)⟧' = (shiftZero A X).Hom ≫ f ≫ (shiftZero A Y).inv := + by + symm + apply nat_iso.naturality_2 +#align category_theory.shift_zero' CategoryTheory.shift_zero' + +variable (C) {A} + +/-- When `i + j = 0`, shifting by `i` and by `j` gives the identity functor -/ +def shiftFunctorCompIsoId (i j : A) (h : i + j = 0) : shiftFunctor C i ⋙ shiftFunctor C j ≅ 𝟭 C := + (shiftFunctorAdd' C i j 0 h).symm ≪≫ shiftFunctorZero C A +#align category_theory.shift_functor_comp_iso_id CategoryTheory.shiftFunctorCompIsoId + +end AddMonoid + +section AddGroup + +variable (C) {A} [AddGroup A] [HasShift C A] + +/-- shifting by `i` and shifting by `j` forms an equivalence when `i + j = 0`. -/ +@[simps] +def shiftEquiv' (i j : A) (h : i + j = 0) : C ≌ C + where + Functor := shiftFunctor C i + inverse := shiftFunctor C j + unitIso := (shiftFunctorCompIsoId C i j h).symm + counitIso := + shiftFunctorCompIsoId C j i (by rw [← add_left_inj j, add_assoc, h, zero_add, add_zero]) + functor_unitIso_comp' X := + by + let E := + equiv_of_tensor_iso_unit (shift_monoidal_functor C A) ⟨i⟩ ⟨j⟩ (eq_to_iso (by ext <;> exact h)) + (eq_to_iso (by ext <;> dsimp <;> rw [← add_left_inj j, add_assoc, h, zero_add, add_zero])) + (Subsingleton.elim _ _) + convert equivalence.functor_unit_iso_comp E X + all_goals + ext X + dsimp [shift_functor_comp_iso_id, unit_of_tensor_iso_unit, shift_functor_add'] + simpa only [eq_to_hom_map, category.assoc] +#align category_theory.shift_equiv' CategoryTheory.shiftEquiv' + +/-- shifting by `n` and shifting by `-n` forms an equivalence. -/ +abbrev shiftEquiv (i : A) : C ≌ C := + shiftEquiv' C i (-i) (add_neg_self i) +#align category_theory.shift_equiv CategoryTheory.shiftEquiv + +variable (X Y : C) (f : X ⟶ Y) + +/-- Shifting by `i` is an equivalence. -/ +instance (i : A) : IsEquivalence (shiftFunctor C i) := + IsEquivalence.ofEquivalence (shiftEquiv C i) + +@[simp] +theorem shiftFunctor_inv (i : A) : (shiftFunctor C i).inv = shiftFunctor C (-i) := + rfl +#align category_theory.shift_functor_inv CategoryTheory.shiftFunctor_inv + +section + +variable (C) + +/-- Shifting by `n` is an essentially surjective functor. -/ +instance shiftFunctor_essSurj (i : A) : EssSurj (shiftFunctor C i) := + Equivalence.essSurj_of_equivalence _ +#align category_theory.shift_functor_ess_surj CategoryTheory.shiftFunctor_essSurj + +end + +variable {C} + +/-- Shifting by `i` and then shifting by `-i` is the identity. -/ +abbrev shiftShiftNeg (i : A) : X⟦i⟧⟦-i⟧ ≅ X := + (shiftEquiv C i).unitIso.symm.app _ +#align category_theory.shift_shift_neg CategoryTheory.shiftShiftNeg + +/-- Shifting by `-i` and then shifting by `i` is the identity. -/ +abbrev shiftNegShift (i : A) : X⟦-i⟧⟦i⟧ ≅ X := + (shiftEquiv C i).counitIso.app _ +#align category_theory.shift_neg_shift CategoryTheory.shiftNegShift + +variable {X Y} + +theorem shift_shift_neg' (i : A) : + f⟦i⟧'⟦-i⟧' = + (shiftFunctorCompIsoId C i (-i) (add_neg_self i)).Hom.app X ≫ + f ≫ (shiftFunctorCompIsoId C i (-i) (add_neg_self i)).inv.app Y := + (NatIso.naturality_2 (shiftFunctorCompIsoId C i (-i) (add_neg_self i)) f).symm +#align category_theory.shift_shift_neg' CategoryTheory.shift_shift_neg' + +theorem shift_neg_shift' (i : A) : + f⟦-i⟧'⟦i⟧' = + (shiftFunctorCompIsoId C (-i) i (neg_add_self i)).Hom.app X ≫ + f ≫ (shiftFunctorCompIsoId C (-i) i (neg_add_self i)).inv.app Y := + (NatIso.naturality_2 (shiftFunctorCompIsoId C (-i) i (neg_add_self i)) f).symm +#align category_theory.shift_neg_shift' CategoryTheory.shift_neg_shift' + +theorem shift_equiv_triangle (n : A) (X : C) : + (shiftShiftNeg X n).inv⟦n⟧' ≫ (shiftNegShift (X⟦n⟧) n).Hom = 𝟙 (X⟦n⟧) := + (shiftEquiv C n).functor_unitIso_comp X +#align category_theory.shift_equiv_triangle CategoryTheory.shift_equiv_triangle + +section + +theorem shift_shiftFunctorCompIsoId_hom_app (n m : A) (h : n + m = 0) (X : C) : + (shiftFunctorCompIsoId C n m h).Hom.app X⟦n⟧' = + (shiftFunctorCompIsoId C m n (by rw [← neg_eq_of_add_eq_zero_left h, add_right_neg])).Hom.app + (X⟦n⟧) := + by + dsimp [shift_functor_comp_iso_id] + simpa only [functor.map_comp, ← shift_functor_add'_zero_add_inv_app n X, ← + shift_functor_add'_add_zero_inv_app n X] using + shift_functor_add'_assoc_inv_app n m n 0 0 n h + (by rw [← neg_eq_of_add_eq_zero_left h, add_right_neg]) (by rw [h, zero_add]) X +#align category_theory.shift_shift_functor_comp_iso_id_hom_app CategoryTheory.shift_shiftFunctorCompIsoId_hom_app + +theorem shift_shiftFunctorCompIsoId_inv_app (n m : A) (h : n + m = 0) (X : C) : + (shiftFunctorCompIsoId C n m h).inv.app X⟦n⟧' = + (shiftFunctorCompIsoId C m n (by rw [← neg_eq_of_add_eq_zero_left h, add_right_neg])).inv.app + (X⟦n⟧) := + by + rw [← cancel_mono ((shift_functor_comp_iso_id C n m h).Hom.app X⟦n⟧'), ← functor.map_comp, + iso.inv_hom_id_app, Functor.map_id, shift_shift_functor_comp_iso_id_hom_app, iso.inv_hom_id_app] + rfl +#align category_theory.shift_shift_functor_comp_iso_id_inv_app CategoryTheory.shift_shiftFunctorCompIsoId_inv_app + +theorem shift_shiftFunctorCompIsoId_add_neg_self_hom_app (n : A) (X : C) : + (shiftFunctorCompIsoId C n (-n) (add_neg_self n)).Hom.app X⟦n⟧' = + (shiftFunctorCompIsoId C (-n) n (neg_add_self n)).Hom.app (X⟦n⟧) := + by apply shift_shift_functor_comp_iso_id_hom_app +#align category_theory.shift_shift_functor_comp_iso_id_add_neg_self_hom_app CategoryTheory.shift_shiftFunctorCompIsoId_add_neg_self_hom_app + +theorem shift_shiftFunctorCompIsoId_add_neg_self_inv_app (n : A) (X : C) : + (shiftFunctorCompIsoId C n (-n) (add_neg_self n)).inv.app X⟦n⟧' = + (shiftFunctorCompIsoId C (-n) n (neg_add_self n)).inv.app (X⟦n⟧) := + by apply shift_shift_functor_comp_iso_id_inv_app +#align category_theory.shift_shift_functor_comp_iso_id_add_neg_self_inv_app CategoryTheory.shift_shiftFunctorCompIsoId_add_neg_self_inv_app + +theorem shift_shiftFunctorCompIsoId_neg_add_self_hom_app (n : A) (X : C) : + (shiftFunctorCompIsoId C (-n) n (neg_add_self n)).Hom.app X⟦-n⟧' = + (shiftFunctorCompIsoId C n (-n) (add_neg_self n)).Hom.app (X⟦-n⟧) := + by apply shift_shift_functor_comp_iso_id_hom_app +#align category_theory.shift_shift_functor_comp_iso_id_neg_add_self_hom_app CategoryTheory.shift_shiftFunctorCompIsoId_neg_add_self_hom_app + +theorem shift_shiftFunctorCompIsoId_neg_add_self_inv_app (n : A) (X : C) : + (shiftFunctorCompIsoId C (-n) n (neg_add_self n)).inv.app X⟦-n⟧' = + (shiftFunctorCompIsoId C n (-n) (add_neg_self n)).inv.app (X⟦-n⟧) := + by apply shift_shift_functor_comp_iso_id_inv_app +#align category_theory.shift_shift_functor_comp_iso_id_neg_add_self_inv_app CategoryTheory.shift_shiftFunctorCompIsoId_neg_add_self_inv_app + +end + +variable {A C} + +open CategoryTheory.Limits + +variable [HasZeroMorphisms C] + +theorem shift_zero_eq_zero (X Y : C) (n : A) : (0 : X ⟶ Y)⟦n⟧' = (0 : X⟦n⟧ ⟶ Y⟦n⟧) := + CategoryTheory.Functor.map_zero _ _ _ +#align category_theory.shift_zero_eq_zero CategoryTheory.shift_zero_eq_zero + +end AddGroup + +section AddCommMonoid + +variable (C) {A} [AddCommMonoid A] [HasShift C A] + +/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ +def shiftFunctorComm (i j : A) : + shiftFunctor C i ⋙ shiftFunctor C j ≅ shiftFunctor C j ⋙ shiftFunctor C i := + (shiftFunctorAdd C i j).symm ≪≫ shiftFunctorAdd' C j i (i + j) (add_comm j i) +#align category_theory.shift_functor_comm CategoryTheory.shiftFunctorComm + +theorem shiftFunctorComm_eq (i j k : A) (h : i + j = k) : + shiftFunctorComm C i j = + (shiftFunctorAdd' C i j k h).symm ≪≫ shiftFunctorAdd' C j i k (by rw [add_comm j i, h]) := + by + subst h + rw [shift_functor_add'_eq_shift_functor_add] + rfl +#align category_theory.shift_functor_comm_eq CategoryTheory.shiftFunctorComm_eq + +theorem shiftFunctorComm_symm (i j : A) : (shiftFunctorComm C i j).symm = shiftFunctorComm C j i := + by + ext1 + dsimp + simpa only [shift_functor_comm_eq C i j (i + j) rfl, + shift_functor_comm_eq C j i (i + j) (add_comm j i)] +#align category_theory.shift_functor_comm_symm CategoryTheory.shiftFunctorComm_symm + +variable {C} (X Y : C) (f : X ⟶ Y) + +/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ +abbrev shiftComm (i j : A) : X⟦i⟧⟦j⟧ ≅ X⟦j⟧⟦i⟧ := + (shiftFunctorComm C i j).app X +#align category_theory.shift_comm CategoryTheory.shiftComm + +@[simp] +theorem shiftComm_symm (i j : A) : (shiftComm X i j).symm = shiftComm X j i := + by + ext1 + exact nat_trans.congr_app (congr_arg iso.hom (shift_functor_comm_symm C i j)) X +#align category_theory.shift_comm_symm CategoryTheory.shiftComm_symm + +variable {X Y} + +/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ +theorem shift_comm' (i j : A) : + f⟦i⟧'⟦j⟧' = (shiftComm _ _ _).Hom ≫ f⟦j⟧'⟦i⟧' ≫ (shiftComm _ _ _).Hom := + by + erw [← shift_comm_symm Y i j, ← (shift_functor_comm C i j).Hom.naturality_assoc f, + iso.hom_inv_id_app, category.comp_id] + rfl +#align category_theory.shift_comm' CategoryTheory.shift_comm' + +@[reassoc.1] +theorem shiftComm_hom_comp (i j : A) : + (shiftComm X i j).Hom ≫ f⟦j⟧'⟦i⟧' = f⟦i⟧'⟦j⟧' ≫ (shiftComm Y i j).Hom := by + rw [shift_comm', ← shift_comm_symm, iso.symm_hom, iso.inv_hom_id_assoc] +#align category_theory.shift_comm_hom_comp CategoryTheory.shiftComm_hom_comp + +theorem shiftFunctorZero_hom_app_shift (n : A) : + (shiftFunctorZero C A).Hom.app (X⟦n⟧) = + (shiftFunctorComm C n 0).Hom.app X ≫ (shiftFunctorZero C A).Hom.app X⟦n⟧' := + by + rw [← shift_functor_add'_zero_add_inv_app n X, shift_functor_comm_eq C n 0 n (add_zero n)] + dsimp + rw [category.assoc, iso.hom_inv_id_app, category.comp_id, shift_functor_add'_add_zero_inv_app] +#align category_theory.shift_functor_zero_hom_app_shift CategoryTheory.shiftFunctorZero_hom_app_shift + +theorem shiftFunctorZero_inv_app_shift (n : A) : + (shiftFunctorZero C A).inv.app (X⟦n⟧) = + (shiftFunctorZero C A).inv.app X⟦n⟧' ≫ (shiftFunctorComm C n 0).inv.app X := + by + rw [← cancel_mono ((shift_functor_zero C A).Hom.app (X⟦n⟧)), category.assoc, iso.inv_hom_id_app, + shift_functor_zero_hom_app_shift, iso.inv_hom_id_app_assoc, ← functor.map_comp, + iso.inv_hom_id_app] + dsimp + rw [Functor.map_id] +#align category_theory.shift_functor_zero_inv_app_shift CategoryTheory.shiftFunctorZero_inv_app_shift + +@[reassoc.1] +theorem shiftFunctorComm_hom_app_comp_shift_shiftFunctorAdd_hom_app (m₁ m₂ m₃ : A) (X : C) : + (shiftFunctorComm C m₁ (m₂ + m₃)).Hom.app X ≫ (shiftFunctorAdd C m₂ m₃).Hom.app X⟦m₁⟧' = + (shiftFunctorAdd C m₂ m₃).Hom.app (X⟦m₁⟧) ≫ + (shiftFunctorComm C m₁ m₂).Hom.app X⟦m₃⟧' ≫ (shiftFunctorComm C m₁ m₃).Hom.app (X⟦m₂⟧) := + by + simp only [← cancel_mono ((shift_functor_comm C m₁ m₃).inv.app (X⟦m₂⟧)), ← + cancel_mono ((shift_functor_comm C m₁ m₂).inv.app X⟦m₃⟧'), category.assoc, iso.hom_inv_id_app] + dsimp + simp only [category.id_comp, ← functor.map_comp, iso.hom_inv_id_app] + dsimp + simp only [Functor.map_id, category.comp_id, shift_functor_comm_eq C _ _ _ rfl, ← + shift_functor_add'_eq_shift_functor_add] + dsimp + simp only [category.assoc, iso.hom_inv_id_app_assoc, iso.inv_hom_id_app_assoc, ← functor.map_comp, + shift_functor_add'_assoc_hom_app_assoc m₂ m₃ m₁ (m₂ + m₃) (m₁ + m₃) (m₁ + (m₂ + m₃)) rfl + (add_comm m₃ m₁) (add_comm _ m₁) X, + ← + shift_functor_add'_assoc_hom_app_assoc m₂ m₁ m₃ (m₁ + m₂) (m₁ + m₃) (m₁ + (m₂ + m₃)) + (add_comm _ _) rfl (by rw [add_comm m₂ m₁, add_assoc]) X, + shift_functor_add'_assoc_hom_app m₁ m₂ m₃ (m₁ + m₂) (m₂ + m₃) (m₁ + (m₂ + m₃)) rfl rfl + (add_assoc _ _ _) X] +#align category_theory.shift_functor_comm_hom_app_comp_shift_shift_functor_add_hom_app CategoryTheory.shiftFunctorComm_hom_app_comp_shift_shiftFunctorAdd_hom_app + +end AddCommMonoid + +variable {C A} {D : Type _} [Category D] [AddMonoid A] [HasShift D A] + +variable (F : C ⥤ D) [Full F] [Faithful F] + +section + +variable (s : A → C ⥤ C) (i : ∀ i, s i ⋙ F ≅ F ⋙ shiftFunctor D i) + +include F s i + +/-- auxiliary definition for `has_shift_of_fully_faithful` -/ +def hasShiftOfFullyFaithfulZero : s 0 ≅ 𝟭 C := + natIsoOfCompFullyFaithful F + (i 0 ≪≫ + isoWhiskerLeft F (shiftFunctorZero D A) ≪≫ + Functor.rightUnitor _ ≪≫ (Functor.leftUnitor _).symm) +#align category_theory.has_shift_of_fully_faithful_zero CategoryTheory.hasShiftOfFullyFaithfulZero + +@[simp] +theorem map_hasShiftOfFullyFaithfulZero_hom_app (X : C) : + F.map ((hasShiftOfFullyFaithfulZero F s i).Hom.app X) = + (i 0).Hom.app X ≫ (shiftFunctorZero D A).Hom.app (F.obj X) := + by + dsimp [has_shift_of_fully_faithful_zero] + simp +#align category_theory.map_has_shift_of_fully_faithful_zero_hom_app CategoryTheory.map_hasShiftOfFullyFaithfulZero_hom_app + +@[simp] +theorem map_hasShiftOfFullyFaithfulZero_inv_app (X : C) : + F.map ((hasShiftOfFullyFaithfulZero F s i).inv.app X) = + (shiftFunctorZero D A).inv.app (F.obj X) ≫ (i 0).inv.app X := + by + dsimp [has_shift_of_fully_faithful_zero] + simp +#align category_theory.map_has_shift_of_fully_faithful_zero_inv_app CategoryTheory.map_hasShiftOfFullyFaithfulZero_inv_app + +/-- auxiliary definition for `has_shift_of_fully_faithful` -/ +def hasShiftOfFullyFaithfulAdd (a b : A) : s (a + b) ≅ s a ⋙ s b := + natIsoOfCompFullyFaithful F + (i (a + b) ≪≫ + isoWhiskerLeft _ (shiftFunctorAdd D a b) ≪≫ + (Functor.associator _ _ _).symm ≪≫ + isoWhiskerRight (i a).symm _ ≪≫ + Functor.associator _ _ _ ≪≫ + isoWhiskerLeft _ (i b).symm ≪≫ (Functor.associator _ _ _).symm) +#align category_theory.has_shift_of_fully_faithful_add CategoryTheory.hasShiftOfFullyFaithfulAdd + +@[simp] +theorem map_hasShiftOfFullyFaithfulAdd_hom_app (a b : A) (X : C) : + F.map ((hasShiftOfFullyFaithfulAdd F s i a b).Hom.app X) = + (i (a + b)).Hom.app X ≫ + (shiftFunctorAdd D a b).Hom.app (F.obj X) ≫ + (i a).inv.app X⟦b⟧' ≫ (i b).inv.app ((s a).obj X) := + by + dsimp [has_shift_of_fully_faithful_add] + simp +#align category_theory.map_has_shift_of_fully_faithful_add_hom_app CategoryTheory.map_hasShiftOfFullyFaithfulAdd_hom_app + +@[simp] +theorem map_hasShiftOfFullyFaithfulAdd_inv_app (a b : A) (X : C) : + F.map ((hasShiftOfFullyFaithfulAdd F s i a b).inv.app X) = + (i b).Hom.app ((s a).obj X) ≫ + (i a).Hom.app X⟦b⟧' ≫ (shiftFunctorAdd D a b).inv.app (F.obj X) ≫ (i (a + b)).inv.app X := + by + dsimp [has_shift_of_fully_faithful_add] + simp +#align category_theory.map_has_shift_of_fully_faithful_add_inv_app CategoryTheory.map_hasShiftOfFullyFaithfulAdd_inv_app + +/-- Given a family of endomorphisms of `C` which are interwined by a fully faithful `F : C ⥤ D` +with shift functors on `D`, we can promote that family to shift functors on `C`. -/ +def hasShiftOfFullyFaithful : HasShift C A := + hasShiftMk C A + { f := s + zero := hasShiftOfFullyFaithfulZero F s i + add := hasShiftOfFullyFaithfulAdd F s i + assoc_hom_app := fun m₁ m₂ m₃ X => + F.map_injective + (by + rw [← cancel_mono ((i m₃).Hom.app ((s m₂).obj ((s m₁).obj X)))] + simp only [functor.map_comp, map_has_shift_of_fully_faithful_add_hom_app, + category.assoc, iso.inv_hom_id_app_assoc, iso.inv_hom_id_app] + erw [(i m₃).Hom.naturality] + have := dcongr_arg (fun a => (i a).Hom.app X) (add_assoc m₁ m₂ m₃) + dsimp at this + dsimp + rw [iso.inv_hom_id_app_assoc, map_has_shift_of_fully_faithful_add_hom_app, this, + eq_to_hom_map, category.comp_id, ← functor.map_comp, category.assoc, + iso.inv_hom_id_app_assoc, functor.map_comp, functor.map_comp, + shift_functor_add_assoc_hom_app_assoc m₁ m₂ m₃ (F.obj X)] + dsimp [shift_functor_add'] + simp only [eq_to_hom_app, category.assoc, eq_to_hom_trans_assoc, eq_to_hom_refl, + category.id_comp, nat_trans.naturality_assoc, functor.comp_map]) + zero_add_hom_app := fun n X => + F.map_injective + (by + have := dcongr_arg (fun a => (i a).Hom.app X) (zero_add n) + dsimp at this + rw [← cancel_mono ((i n).Hom.app ((s 0).obj X))] + simp only [this, map_has_shift_of_fully_faithful_add_hom_app, + shift_functor_add_zero_add_hom_app, eq_to_hom_map, category.assoc, + eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp, iso.inv_hom_id_app, + functor.map_comp] + erw [(i n).Hom.naturality] + dsimp + simp) + add_zero_hom_app := fun n X => + F.map_injective + (by + have := dcongr_arg (fun a => (i a).Hom.app X) (add_zero n) + dsimp at this + simpa [this, ← nat_trans.naturality_assoc, eq_to_hom_map, + shift_functor_add_add_zero_hom_app] ) } +#align category_theory.has_shift_of_fully_faithful CategoryTheory.hasShiftOfFullyFaithful + +end + +end CategoryTheory + diff --git a/Mathbin/CategoryTheory/Triangulated/Basic.lean b/Mathbin/CategoryTheory/Triangulated/Basic.lean index 661f2dfed0..934518f868 100644 --- a/Mathbin/CategoryTheory/Triangulated/Basic.lean +++ b/Mathbin/CategoryTheory/Triangulated/Basic.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Luke Kershaw ! This file was ported from Lean 3 source module category_theory.triangulated.basic -! leanprover-community/mathlib commit f7707875544ef1f81b32cb68c79e0e24e45a0e76 +! leanprover-community/mathlib commit 6876fa15e3158ff3e4a4e2af1fb6e1945c6e8803 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ import Mathbin.Data.Int.Basic -import Mathbin.CategoryTheory.Shift +import Mathbin.CategoryTheory.Shift.Basic /-! # Triangles @@ -153,5 +153,39 @@ instance triangleCategory : Category (Triangle C) comp A B C f g := f.comp g #align category_theory.pretriangulated.triangle_category CategoryTheory.Pretriangulated.triangleCategory +/-- a constructor for morphisms of triangles -/ +@[simps] +def Triangle.homMk (A B : Triangle C) (hom₁ : A.obj₁ ⟶ B.obj₁) (hom₂ : A.obj₂ ⟶ B.obj₂) + (hom₃ : A.obj₃ ⟶ B.obj₃) (comm₁ : A.mor₁ ≫ hom₂ = hom₁ ≫ B.mor₁) + (comm₂ : A.mor₂ ≫ hom₃ = hom₂ ≫ B.mor₂) (comm₃ : A.mor₃ ≫ hom₁⟦1⟧' = hom₃ ≫ B.mor₃) : A ⟶ B := + { hom₁ + hom₂ + hom₃ + comm₁' := comm₁ + comm₂' := comm₂ + comm₃' := comm₃ } +#align category_theory.pretriangulated.triangle.hom_mk CategoryTheory.Pretriangulated.Triangle.homMk + +/-- a constructor for isomorphisms of triangles -/ +@[simps] +def Triangle.isoMk (A B : Triangle C) (iso₁ : A.obj₁ ≅ B.obj₁) (iso₂ : A.obj₂ ≅ B.obj₂) + (iso₃ : A.obj₃ ≅ B.obj₃) (comm₁ : A.mor₁ ≫ iso₂.Hom = iso₁.Hom ≫ B.mor₁) + (comm₂ : A.mor₂ ≫ iso₃.Hom = iso₂.Hom ≫ B.mor₂) + (comm₃ : A.mor₃ ≫ iso₁.Hom⟦1⟧' = iso₃.Hom ≫ B.mor₃) : A ≅ B + where + Hom := Triangle.homMk _ _ iso₁.Hom iso₂.Hom iso₃.Hom comm₁ comm₂ comm₃ + inv := + Triangle.homMk _ _ iso₁.inv iso₂.inv iso₃.inv + (by + simp only [← cancel_mono iso₂.hom, assoc, iso.inv_hom_id, comp_id, comm₁, + iso.inv_hom_id_assoc]) + (by + simp only [← cancel_mono iso₃.hom, assoc, iso.inv_hom_id, comp_id, comm₂, + iso.inv_hom_id_assoc]) + (by + simp only [← cancel_mono (iso₁.hom⟦(1 : ℤ)⟧'), assoc, ← functor.map_comp, iso.inv_hom_id, + CategoryTheory.Functor.map_id, comp_id, comm₃, iso.inv_hom_id_assoc]) +#align category_theory.pretriangulated.triangle.iso_mk CategoryTheory.Pretriangulated.Triangle.isoMk + end CategoryTheory.Pretriangulated diff --git a/Mathbin/CategoryTheory/Triangulated/Pretriangulated.lean b/Mathbin/CategoryTheory/Triangulated/Pretriangulated.lean index 767a6f5af8..a98f83ed62 100644 --- a/Mathbin/CategoryTheory/Triangulated/Pretriangulated.lean +++ b/Mathbin/CategoryTheory/Triangulated/Pretriangulated.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Luke Kershaw ! This file was ported from Lean 3 source module category_theory.triangulated.pretriangulated -! leanprover-community/mathlib commit 19786714ebe478f40b503acb4705fb058ba47303 +! leanprover-community/mathlib commit 6876fa15e3158ff3e4a4e2af1fb6e1945c6e8803 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ import Mathbin.CategoryTheory.Preadditive.AdditiveFunctor -import Mathbin.CategoryTheory.Shift +import Mathbin.CategoryTheory.Shift.Basic import Mathbin.CategoryTheory.Triangulated.Rotate /-! @@ -163,98 +163,6 @@ theorem comp_dist_triangle_mor_zero₃₁ (T) (_ : T ∈ (dist_triang C)) : TODO: If `C` is pretriangulated with respect to a shift, then `Cᵒᵖ` is pretriangulated with respect to the inverse shift. -/ -omit hC - -/-- -The underlying structure of a triangulated functor between pretriangulated categories `C` and `D` -is a functor `F : C ⥤ D` together with given functorial isomorphisms `ξ X : F(X⟦1⟧) ⟶ F(X)⟦1⟧`. --/ -structure TriangulatedFunctorStruct extends C ⥤ D where - commShift : shiftFunctor C (1 : ℤ) ⋙ to_functor ≅ to_functor ⋙ shiftFunctor D (1 : ℤ) -#align category_theory.pretriangulated.triangulated_functor_struct CategoryTheory.Pretriangulated.TriangulatedFunctorStruct - -namespace TriangulatedFunctorStruct - -/-- The identity `triangulated_functor_struct`. -/ -def id : TriangulatedFunctorStruct C C where - obj X := X - map _ _ f := f - commShift := by rfl -#align category_theory.pretriangulated.triangulated_functor_struct.id CategoryTheory.Pretriangulated.TriangulatedFunctorStruct.id - -instance : Inhabited (TriangulatedFunctorStruct C C) := - ⟨id C⟩ - -variable {C D} - -/-- Given a `triangulated_functor_struct` we can define a functor from triangles of `C` to -triangles of `D`. --/ -@[simps] -def mapTriangle (F : TriangulatedFunctorStruct C D) : Triangle C ⥤ Triangle D - where - obj T := Triangle.mk (F.map T.mor₁) (F.map T.mor₂) (F.map T.mor₃ ≫ F.commShift.Hom.app T.obj₁) - map S T f := - { hom₁ := F.map f.hom₁ - hom₂ := F.map f.hom₂ - hom₃ := F.map f.hom₃ - comm₁' := by - dsimp - simp only [← F.to_functor.map_comp, f.comm₁] - comm₂' := by - dsimp - simp only [← F.to_functor.map_comp, f.comm₂] - comm₃' := by - dsimp - erw [category.assoc, ← F.comm_shift.hom.naturality] - simp only [functor.comp_map, ← F.to_functor.map_comp_assoc, f.comm₃] } -#align category_theory.pretriangulated.triangulated_functor_struct.map_triangle CategoryTheory.Pretriangulated.TriangulatedFunctorStruct.mapTriangle - -end TriangulatedFunctorStruct - -include hC - -variable (C D) [Pretriangulated D] - -/-- A triangulated functor between pretriangulated categories `C` and `D` is a functor `F : C ⥤ D` -together with given functorial isomorphisms `ξ X : F(X⟦1⟧) ⟶ F(X)⟦1⟧` such that for every -distinguished triangle `(X,Y,Z,f,g,h)` of `C`, the triangle -`(F(X), F(Y), F(Z), F(f), F(g), F(h) ≫ (ξ X))` is a distinguished triangle of `D`. -See --/ -structure TriangulatedFunctor extends TriangulatedFunctorStruct C D where - map_distinguished' : - ∀ T : Triangle C, - T ∈ (dist_triang C) → to_triangulated_functor_struct.mapTriangle.obj T ∈ (dist_triang D) -#align category_theory.pretriangulated.triangulated_functor CategoryTheory.Pretriangulated.TriangulatedFunctor - -instance : Inhabited (TriangulatedFunctor C C) := - ⟨{ obj := fun X => X - map := fun _ _ f => f - commShift := by rfl - map_distinguished' := by - rintro ⟨_, _, _, _⟩ Tdt - dsimp at * - rwa [category.comp_id] }⟩ - -variable {C D} - -/-- -Given a `triangulated_functor` we can define a functor from triangles of `C` to triangles of `D`. --/ -@[simps] -def TriangulatedFunctor.mapTriangle (F : TriangulatedFunctor C D) : Triangle C ⥤ Triangle D := - F.toTriangulatedFunctorStruct.mapTriangle -#align category_theory.pretriangulated.triangulated_functor.map_triangle CategoryTheory.Pretriangulated.TriangulatedFunctor.mapTriangle - -/-- Given a `triangulated_functor` and a distinguished triangle `T` of `C`, then the triangle it -maps onto in `D` is also distinguished. --/ -theorem TriangulatedFunctor.map_distinguished (F : TriangulatedFunctor C D) (T : Triangle C) - (h : T ∈ (dist_triang C)) : F.mapTriangle.obj T ∈ (dist_triang D) := - F.map_distinguished' T h -#align category_theory.pretriangulated.triangulated_functor.map_distinguished CategoryTheory.Pretriangulated.TriangulatedFunctor.map_distinguished - end Pretriangulated end CategoryTheory diff --git a/Mathbin/CategoryTheory/Triangulated/Rotate.lean b/Mathbin/CategoryTheory/Triangulated/Rotate.lean index eea420940c..05356d4432 100644 --- a/Mathbin/CategoryTheory/Triangulated/Rotate.lean +++ b/Mathbin/CategoryTheory/Triangulated/Rotate.lean @@ -1,10 +1,10 @@ /- Copyright (c) 2021 Luke Kershaw. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Luke Kershaw +Authors: Luke Kershaw, Joël Riou ! This file was ported from Lean 3 source module category_theory.triangulated.rotate -! leanprover-community/mathlib commit 19786714ebe478f40b503acb4705fb058ba47303 +! leanprover-community/mathlib commit 6876fa15e3158ff3e4a4e2af1fb6e1945c6e8803 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -59,8 +59,6 @@ def Triangle.rotate (T : Triangle C) : Triangle C := section -attribute [local semireducible] shift_shift_neg shift_neg_shift - /-- Given a triangle of the form: ``` f g h @@ -82,92 +80,6 @@ def Triangle.invRotate (T : Triangle C) : Triangle C := end -namespace TriangleMorphism - -variable {T₁ T₂ T₃ T₄ : Triangle C} - -open Triangle - -/-- You can also rotate a triangle morphism to get a morphism between the two rotated triangles. -Given a triangle morphism of the form: -``` - f g h - X ───> Y ───> Z ───> X⟦1⟧ - │ │ │ │ - │a │b │c │a⟦1⟧ - V V V V - X' ───> Y' ───> Z' ───> X'⟦1⟧ - f' g' h' -``` -applying `rotate` gives a triangle morphism of the form: - -``` - g h -f⟦1⟧ - Y ───> Z ───> X⟦1⟧ ───> Y⟦1⟧ - │ │ │ │ - │b │c │a⟦1⟧ │b⟦1⟧' - V V V V - Y' ───> Z' ───> X'⟦1⟧ ───> Y'⟦1⟧ - g' h' -f'⟦1⟧ -``` --/ -@[simps] -def rotate (f : TriangleMorphism T₁ T₂) : TriangleMorphism T₁.rotate T₂.rotate - where - hom₁ := f.hom₂ - hom₂ := f.hom₃ - hom₃ := f.hom₁⟦1⟧' - comm₃' := by - dsimp - simp only [rotate_mor₃, comp_neg, neg_comp, ← functor.map_comp, f.comm₁] -#align category_theory.pretriangulated.triangle_morphism.rotate CategoryTheory.Pretriangulated.TriangleMorphism.rotate - -/-- Given a triangle morphism of the form: -``` - f g h - X ───> Y ───> Z ───> X⟦1⟧ - │ │ │ │ - │a │b │c │a⟦1⟧ - V V V V - X' ───> Y' ───> Z' ───> X'⟦1⟧ - f' g' h' -``` -applying `inv_rotate` gives a triangle morphism that can be thought of as: -``` - -h⟦-1⟧ f g - Z⟦-1⟧ ───> X ───> Y ───> Z - │ │ │ │ - │c⟦-1⟧' │a │b │c - V V V V - Z'⟦-1⟧ ───> X' ───> Y' ───> Z' - -h'⟦-1⟧ f' g' -``` -(note that this diagram doesn't technically fit the definition of triangle morphism, -as `Z⟦-1⟧⟦1⟧` is not necessarily equal to `Z`, and `Z'⟦-1⟧⟦1⟧` is not necessarily equal to `Z'`, -but they are isomorphic, by the `counit_iso` of `shift C`) --/ -@[simps] -def invRotate (f : TriangleMorphism T₁ T₂) : TriangleMorphism T₁.invRotate T₂.invRotate - where - hom₁ := f.hom₃⟦-1⟧' - hom₂ := f.hom₁ - hom₃ := f.hom₂ - comm₁' := by - dsimp [inv_rotate_mor₁] - simp only [discrete.functor_map_id, id_comp, preadditive.comp_neg, assoc, neg_inj, - nat_trans.id_app, preadditive.neg_comp] - rw [← functor.map_comp_assoc, ← f.comm₃, functor.map_comp_assoc, μ_naturality_assoc, - nat_trans.naturality, functor.id_map] - comm₃' := by - dsimp - simp only [discrete.functor_map_id, id_comp, μ_inv_naturality, category.assoc, nat_trans.id_app, - unit_of_tensor_iso_unit_inv_app] - erw [ε_naturality_assoc] - rw [comm₂_assoc] -#align category_theory.pretriangulated.triangle_morphism.inv_rotate CategoryTheory.Pretriangulated.TriangleMorphism.invRotate - -end TriangleMorphism - variable (C) /-- Rotating triangles gives an endofunctor on the category of triangles in `C`. @@ -176,165 +88,67 @@ variable (C) def rotate : Triangle C ⥤ Triangle C where obj := Triangle.rotate - map _ _ f := f.rotate + map T₁ T₂ f := + { hom₁ := f.hom₂ + hom₂ := f.hom₃ + hom₃ := f.hom₁⟦1⟧' + comm₃' := by + dsimp + simp only [comp_neg, neg_comp, ← functor.map_comp, f.comm₁] } #align category_theory.pretriangulated.rotate CategoryTheory.Pretriangulated.rotate +example : ℕ := + 42 + /-- The inverse rotation of triangles gives an endofunctor on the category of triangles in `C`. -/ @[simps] def invRotate : Triangle C ⥤ Triangle C where obj := Triangle.invRotate - map _ _ f := f.invRotate + map T₁ T₂ f := + { hom₁ := f.hom₃⟦-1⟧' + hom₂ := f.hom₁ + hom₃ := f.hom₂ + comm₁' := by + dsimp + rw [neg_comp, assoc, comp_neg, neg_inj, ← functor.map_comp_assoc, ← f.comm₃, + functor.map_comp, assoc] + erw [← nat_trans.naturality] + rfl + comm₃' := by + dsimp + erw [← f.comm₂_assoc, assoc, ← nat_trans.naturality] + rfl } #align category_theory.pretriangulated.inv_rotate CategoryTheory.Pretriangulated.invRotate variable {C} variable [∀ n : ℤ, Functor.Additive (shiftFunctor C n)] -/-- There is a natural map from a triangle to the `inv_rotate` of its `rotate`. -/ -@[simps] -def toInvRotateRotate (T : Triangle C) : T ⟶ (invRotate C).obj ((rotate C).obj T) - where - hom₁ := (shiftShiftNeg _ _).inv - hom₂ := 𝟙 T.obj₂ - hom₃ := 𝟙 T.obj₃ - comm₃' := by - dsimp - simp only [ε_app_obj, eq_to_iso.hom, discrete.functor_map_id, id_comp, eq_to_iso.inv, - category.assoc, obj_μ_inv_app, functor.map_comp, nat_trans.id_app, obj_ε_app, - unit_of_tensor_iso_unit_inv_app] - erw [μ_inv_hom_app_assoc] - rfl -#align category_theory.pretriangulated.to_inv_rotate_rotate CategoryTheory.Pretriangulated.toInvRotateRotate - -/-- There is a natural transformation between the identity functor on triangles in `C`, -and the composition of a rotation with an inverse rotation. --/ -@[simps] -def rotCompInvRotHom : 𝟭 (Triangle C) ⟶ rotate C ⋙ invRotate C - where - app := toInvRotateRotate - naturality' := by - introv ; ext - · dsimp - simp only [nat_iso.cancel_nat_iso_inv_right_assoc, discrete.functor_map_id, id_comp, - μ_inv_naturality, assoc, nat_trans.id_app, unit_of_tensor_iso_unit_inv_app] - erw [ε_naturality] - · dsimp - rw [comp_id, id_comp] - · dsimp - rw [comp_id, id_comp] -#align category_theory.pretriangulated.rot_comp_inv_rot_hom CategoryTheory.Pretriangulated.rotCompInvRotHom - -/-- There is a natural map from the `inv_rotate` of the `rotate` of a triangle to itself. -/ -@[simps] -def fromInvRotateRotate (T : Triangle C) : (invRotate C).obj ((rotate C).obj T) ⟶ T - where - hom₁ := (shiftEquiv C 1).unitInv.app T.obj₁ - hom₂ := 𝟙 T.obj₂ - hom₃ := 𝟙 T.obj₃ - comm₃' := by - dsimp - rw [unit_of_tensor_iso_unit_inv_app, ε_app_obj] - simp only [discrete.functor_map_id, nat_trans.id_app, id_comp, assoc, functor.map_comp, - obj_μ_app, obj_ε_inv_app, comp_id, μ_inv_hom_app_assoc] - erw [μ_inv_hom_app, μ_inv_hom_app_assoc, category.comp_id] -#align category_theory.pretriangulated.from_inv_rotate_rotate CategoryTheory.Pretriangulated.fromInvRotateRotate - -/-- There is a natural transformation between the composition of a rotation with an inverse rotation -on triangles in `C`, and the identity functor. --/ -@[simps] -def rotCompInvRotInv : rotate C ⋙ invRotate C ⟶ 𝟭 (Triangle C) where app := fromInvRotateRotate -#align category_theory.pretriangulated.rot_comp_inv_rot_inv CategoryTheory.Pretriangulated.rotCompInvRotInv +attribute [local simp] + shift_shift_neg' shift_neg_shift' shift_shift_functor_comp_iso_id_add_neg_self_inv_app shift_shift_functor_comp_iso_id_add_neg_self_hom_app -/-- The natural transformations between the identity functor on triangles in `C` and the composition -of a rotation with an inverse rotation are natural isomorphisms (they are isomorphisms in the -category of functors). --/ +/-- The unit isomorphism of the auto-equivalence of categories `triangle_rotation C` of +`triangle C` given by the rotation of triangles. -/ @[simps] -def rotCompInvRot : 𝟭 (Triangle C) ≅ rotate C ⋙ invRotate C - where - Hom := rotCompInvRotHom - inv := rotCompInvRotInv +def rotCompInvRot : 𝟭 (Triangle C) ≅ rotate C ⋙ invRotate C := + NatIso.ofComponents + (fun T => + Triangle.isoMk _ _ ((shiftEquiv C (1 : ℤ)).unitIso.app T.obj₁) (Iso.refl _) (Iso.refl _) + (by tidy) (by tidy) (by tidy)) + (by tidy) #align category_theory.pretriangulated.rot_comp_inv_rot CategoryTheory.Pretriangulated.rotCompInvRot -/-- There is a natural map from the `rotate` of the `inv_rotate` of a triangle to itself. -/ +/-- The counit isomorphism of the auto-equivalence of categories `triangle_rotation C` of +`triangle C` given by the rotation of triangles. -/ @[simps] -def fromRotateInvRotate (T : Triangle C) : (rotate C).obj ((invRotate C).obj T) ⟶ T - where - hom₁ := 𝟙 T.obj₁ - hom₂ := 𝟙 T.obj₂ - hom₃ := (shiftEquiv C 1).counit.app T.obj₃ - comm₂' := by - dsimp - rw [unit_of_tensor_iso_unit_inv_app] - simp only [discrete.functor_map_id, nat_trans.id_app, id_comp, add_neg_equiv_counit_iso_hom, - eq_to_hom_refl, nat_trans.comp_app, assoc, μ_inv_hom_app_assoc, ε_hom_inv_app] - exact category.comp_id _ - comm₃' := by - dsimp - simp only [discrete.functor_map_id, nat_trans.id_app, id_comp, functor.map_neg, - functor.map_comp, obj_μ_app, obj_ε_inv_app, comp_id, assoc, μ_naturality_assoc, neg_neg, - CategoryTheory.Functor.map_id, add_neg_equiv_counit_iso_hom, eq_to_hom_refl, - nat_trans.comp_app] - erw [μ_inv_hom_app, category.comp_id, obj_zero_map_μ_app] - rw [discrete.functor_map_id, nat_trans.id_app, comp_id] -#align category_theory.pretriangulated.from_rotate_inv_rotate CategoryTheory.Pretriangulated.fromRotateInvRotate - -/-- There is a natural transformation between the composition of an inverse rotation with a rotation -on triangles in `C`, and the identity functor. --/ -@[simps] -def invRotCompRotHom : invRotate C ⋙ rotate C ⟶ 𝟭 (Triangle C) where app := fromRotateInvRotate -#align category_theory.pretriangulated.inv_rot_comp_rot_hom CategoryTheory.Pretriangulated.invRotCompRotHom - -/-- There is a natural map from a triangle to the `rotate` of its `inv_rotate`. -/ -@[simps] -def toRotateInvRotate (T : Triangle C) : T ⟶ (rotate C).obj ((invRotate C).obj T) - where - hom₁ := 𝟙 T.obj₁ - hom₂ := 𝟙 T.obj₂ - hom₃ := (shiftEquiv C 1).counitInv.app T.obj₃ - comm₃' := by - dsimp - rw [CategoryTheory.Functor.map_id] - simp only [comp_id, add_neg_equiv_counit_iso_inv, eq_to_hom_refl, id_comp, nat_trans.comp_app, - discrete.functor_map_id, nat_trans.id_app, functor.map_neg, functor.map_comp, obj_μ_app, - obj_ε_inv_app, assoc, μ_naturality_assoc, neg_neg, μ_inv_hom_app_assoc] - erw [μ_inv_hom_app, category.comp_id, obj_zero_map_μ_app] - simp only [discrete.functor_map_id, nat_trans.id_app, comp_id, ε_hom_inv_app_assoc] -#align category_theory.pretriangulated.to_rotate_inv_rotate CategoryTheory.Pretriangulated.toRotateInvRotate - -/-- There is a natural transformation between the identity functor on triangles in `C`, -and the composition of an inverse rotation with a rotation. --/ -@[simps] -def invRotCompRotInv : 𝟭 (Triangle C) ⟶ invRotate C ⋙ rotate C - where - app := toRotateInvRotate - naturality' := by - introv ; ext - · dsimp - rw [comp_id, id_comp] - · dsimp - rw [comp_id, id_comp] - · dsimp - rw [add_neg_equiv_counit_iso_inv, eq_to_hom_map, eq_to_hom_refl, id_comp] - simp only [nat_trans.comp_app, assoc] - erw [μ_inv_naturality, ε_naturality_assoc] -#align category_theory.pretriangulated.inv_rot_comp_rot_inv CategoryTheory.Pretriangulated.invRotCompRotInv - -/-- The natural transformations between the composition of a rotation with an inverse rotation -on triangles in `C`, and the identity functor on triangles are natural isomorphisms -(they are isomorphisms in the category of functors). --/ -@[simps] -def invRotCompRot : invRotate C ⋙ rotate C ≅ 𝟭 (Triangle C) - where - Hom := invRotCompRotHom - inv := invRotCompRotInv +def invRotCompRot : invRotate C ⋙ rotate C ≅ 𝟭 (Triangle C) := + NatIso.ofComponents + (fun T => + Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) ((shiftEquiv C (1 : ℤ)).counitIso.app T.obj₃) + (by tidy) (by tidy) (by tidy)) + (by tidy) #align category_theory.pretriangulated.inv_rot_comp_rot CategoryTheory.Pretriangulated.invRotCompRot variable (C) @@ -348,19 +162,6 @@ def triangleRotation : Equivalence (Triangle C) (Triangle C) inverse := invRotate C unitIso := rotCompInvRot counitIso := invRotCompRot - functor_unitIso_comp' := by - introv ; ext - · dsimp - rw [comp_id] - · dsimp - rw [comp_id] - · dsimp - rw [unit_of_tensor_iso_unit_inv_app] - simp only [discrete.functor_map_id, nat_trans.id_app, id_comp, functor.map_comp, obj_ε_app, - obj_μ_inv_app, assoc, add_neg_equiv_counit_iso_hom, eq_to_hom_refl, nat_trans.comp_app, - ε_inv_app_obj, comp_id, μ_inv_hom_app_assoc] - erw [μ_inv_hom_app_assoc, μ_inv_hom_app] - rfl #align category_theory.pretriangulated.triangle_rotation CategoryTheory.Pretriangulated.triangleRotation variable {C} diff --git a/Mathbin/Combinatorics/SimpleGraph/Connectivity.lean b/Mathbin/Combinatorics/SimpleGraph/Connectivity.lean index 1cbfe0c73d..8a78e7dbea 100644 --- a/Mathbin/Combinatorics/SimpleGraph/Connectivity.lean +++ b/Mathbin/Combinatorics/SimpleGraph/Connectivity.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller ! This file was ported from Lean 3 source module combinatorics.simple_graph.connectivity -! leanprover-community/mathlib commit 832f7b9162039c28b9361289c8681f155cae758f +! leanprover-community/mathlib commit b99e2d58a5e6861833fa8de11e51a81144258db4 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -2582,6 +2582,16 @@ protected theorem Reachable.map {G : SimpleGraph V} {G' : SimpleGraph V'} (f : G h.elim fun p => ⟨p.map f⟩ #align simple_graph.reachable.map SimpleGraph.Reachable.map +theorem Iso.reachable_iff {G : SimpleGraph V} {G' : SimpleGraph V'} {φ : G ≃g G'} {u v : V} : + G'.Reachable (φ u) (φ v) ↔ G.Reachable u v := + ⟨fun r => φ.left_inv u ▸ φ.left_inv v ▸ r.map φ.symm.toHom, Reachable.map φ.toHom⟩ +#align simple_graph.iso.reachable_iff SimpleGraph.Iso.reachable_iff + +theorem Iso.symm_apply_reachable {G : SimpleGraph V} {G' : SimpleGraph V'} {φ : G ≃g G'} {u : V} + {v : V'} : G.Reachable (φ.symm v) u ↔ G'.Reachable v (φ u) := by + rw [← iso.reachable_iff, RelIso.apply_symm_apply] +#align simple_graph.iso.symm_apply_reachable SimpleGraph.Iso.symm_apply_reachable + variable (G) #print SimpleGraph.reachable_is_equivalence /- @@ -2716,6 +2726,11 @@ protected theorem eq {v w : V} : #align simple_graph.connected_component.eq SimpleGraph.ConnectedComponent.eq -/ +theorem connectedComponentMk_eq_of_adj {v w : V} (a : G.Adj v w) : + G.connectedComponentMk v = G.connectedComponentMk w := + ConnectedComponent.sound a.Reachable +#align simple_graph.connected_component.connected_component_mk_eq_of_adj SimpleGraph.ConnectedComponent.connectedComponentMk_eq_of_adj + #print SimpleGraph.ConnectedComponent.lift /- /-- The `connected_component` specialization of `quot.lift`. Provides the stronger assumption that the vertices are connected by a path. -/ @@ -2809,6 +2824,120 @@ theorem map_comp (C : G.ConnectedComponent) (φ : G →g G') (ψ : G' →g G'') exact fun _ => rfl #align simple_graph.connected_component.map_comp SimpleGraph.ConnectedComponent.map_comp +variable {φ : G ≃g G'} {v : V} {v' : V'} + +@[simp] +theorem iso_image_comp_eq_map_iff_eq_comp {C : G.ConnectedComponent} : + G'.connectedComponentMk (φ v) = C.map ↑(↑φ : G ↪g G') ↔ G.connectedComponentMk v = C := + by + refine' C.ind fun u => _ + simp only [iso.reachable_iff, connected_component.map_mk, RelEmbedding.coe_coeFn, + RelIso.coe_coeFn, connected_component.eq] +#align simple_graph.connected_component.iso_image_comp_eq_map_iff_eq_comp SimpleGraph.ConnectedComponent.iso_image_comp_eq_map_iff_eq_comp + +@[simp] +theorem iso_inv_image_comp_eq_iff_eq_map {C : G.ConnectedComponent} : + G.connectedComponentMk (φ.symm v') = C ↔ G'.connectedComponentMk v' = C.map φ := + by + refine' C.ind fun u => _ + simp only [iso.symm_apply_reachable, connected_component.eq, coe_coe, connected_component.map_mk, + RelEmbedding.coe_coeFn, RelIso.coe_coeFn] +#align simple_graph.connected_component.iso_inv_image_comp_eq_iff_eq_map SimpleGraph.ConnectedComponent.iso_inv_image_comp_eq_iff_eq_map + +end connectedComponent + +namespace Iso + +/-- An isomorphism of graphs induces a bijection of connected components. -/ +@[simps] +def connectedComponentEquiv (φ : G ≃g G') : G.ConnectedComponent ≃ G'.ConnectedComponent + where + toFun := ConnectedComponent.map φ + invFun := ConnectedComponent.map φ.symm + left_inv C := + ConnectedComponent.ind (fun v => congr_arg G.connectedComponentMk (Equiv.left_inv φ.toEquiv v)) + C + right_inv C := + ConnectedComponent.ind + (fun v => congr_arg G'.connectedComponentMk (Equiv.right_inv φ.toEquiv v)) C +#align simple_graph.iso.connected_component_equiv SimpleGraph.Iso.connectedComponentEquiv + +@[simp] +theorem connectedComponentEquiv_refl : (Iso.refl : G ≃g G).connectedComponentEquiv = Equiv.refl _ := + by + ext ⟨v⟩ + rfl +#align simple_graph.iso.connected_component_equiv_refl SimpleGraph.Iso.connectedComponentEquiv_refl + +@[simp] +theorem connectedComponentEquiv_symm (φ : G ≃g G') : + φ.symm.connectedComponentEquiv = φ.connectedComponentEquiv.symm := + by + ext ⟨_⟩ + rfl +#align simple_graph.iso.connected_component_equiv_symm SimpleGraph.Iso.connectedComponentEquiv_symm + +@[simp] +theorem connectedComponentEquiv_trans (φ : G ≃g G') (φ' : G' ≃g G'') : + connectedComponentEquiv (φ.trans φ') = + φ.connectedComponentEquiv.trans φ'.connectedComponentEquiv := + by + ext ⟨_⟩ + rfl +#align simple_graph.iso.connected_component_equiv_trans SimpleGraph.Iso.connectedComponentEquiv_trans + +end Iso + +namespace connectedComponent + +/-- The set of vertices in a connected component of a graph. -/ +def supp (C : G.ConnectedComponent) := + { v | G.connectedComponentMk v = C } +#align simple_graph.connected_component.supp SimpleGraph.ConnectedComponent.supp + +@[ext] +theorem supp_injective : + Function.Injective (ConnectedComponent.supp : G.ConnectedComponent → Set V) := + by + refine' connected_component.ind₂ _ + intro v w + simp only [connected_component.supp, Set.ext_iff, connected_component.eq, Set.mem_setOf_eq] + intro h + rw [reachable_comm, h] +#align simple_graph.connected_component.supp_injective SimpleGraph.ConnectedComponent.supp_injective + +@[simp] +theorem supp_inj {C D : G.ConnectedComponent} : C.supp = D.supp ↔ C = D := + ConnectedComponent.supp_injective.eq_iff +#align simple_graph.connected_component.supp_inj SimpleGraph.ConnectedComponent.supp_inj + +instance : SetLike G.ConnectedComponent V + where + coe := ConnectedComponent.supp + coe_injective' := ConnectedComponent.supp_injective + +@[simp] +theorem mem_supp_iff (C : G.ConnectedComponent) (v : V) : + v ∈ C.supp ↔ G.connectedComponentMk v = C := + Iff.rfl +#align simple_graph.connected_component.mem_supp_iff SimpleGraph.ConnectedComponent.mem_supp_iff + +theorem connectedComponentMk_mem {v : V} : v ∈ G.connectedComponentMk v := + rfl +#align simple_graph.connected_component.connected_component_mk_mem SimpleGraph.ConnectedComponent.connectedComponentMk_mem + +/-- The equivalence between connected components, induced by an isomorphism of graphs, +itself defines an equivalence on the supports of each connected component. +-/ +def isoEquivSupp (φ : G ≃g G') (C : G.ConnectedComponent) : + C.supp ≃ (φ.connectedComponentEquiv C).supp + where + toFun v := ⟨φ v, ConnectedComponent.iso_image_comp_eq_map_iff_eq_comp.mpr v.Prop⟩ + invFun v' := ⟨φ.symm v', ConnectedComponent.iso_inv_image_comp_eq_iff_eq_map.mpr v'.Prop⟩ + left_inv v := Subtype.ext_val (φ.toEquiv.left_inv ↑v) + right_inv v := Subtype.ext_val (φ.toEquiv.right_inv ↑v) +#align simple_graph.connected_component.iso_equiv_supp SimpleGraph.ConnectedComponent.isoEquivSupp + end connectedComponent #print SimpleGraph.Subgraph.Connected /- diff --git a/Mathbin/Combinatorics/SimpleGraph/Ends/Defs.lean b/Mathbin/Combinatorics/SimpleGraph/Ends/Defs.lean index 8d4877d977..0d8b7093ca 100644 --- a/Mathbin/Combinatorics/SimpleGraph/Ends/Defs.lean +++ b/Mathbin/Combinatorics/SimpleGraph/Ends/Defs.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anand Rao, Rémi Bottinelli ! This file was ported from Lean 3 source module combinatorics.simple_graph.ends.defs -! leanprover-community/mathlib commit 8195826f5c428fc283510bc67303dd4472d78498 +! leanprover-community/mathlib commit b99e2d58a5e6861833fa8de11e51a81144258db4 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -59,10 +59,11 @@ theorem ComponentCompl.supp_inj {C D : G.ComponentCompl K} : C.supp = D.supp ↔ ComponentCompl.supp_injective.eq_iff #align simple_graph.component_compl.supp_inj SimpleGraph.ComponentCompl.supp_inj -instance : SetLike (G.ComponentCompl K) V +instance ComponentCompl.setLike : SetLike (G.ComponentCompl K) V where coe := ComponentCompl.supp coe_injective' C D := ComponentCompl.supp_inj.mp +#align simple_graph.component_compl.set_like SimpleGraph.ComponentCompl.setLike @[simp] theorem ComponentCompl.mem_supp_iff {v : V} {C : ComponentCompl G K} : diff --git a/Mathbin/Data/Finset/Basic.lean b/Mathbin/Data/Finset/Basic.lean index d2557f5687..298cf30431 100644 --- a/Mathbin/Data/Finset/Basic.lean +++ b/Mathbin/Data/Finset/Basic.lean @@ -4980,9 +4980,9 @@ theorem singleton_disjUnionᵢ (a : α) {h} : Finset.disjUnion {a} t h = t a := /- warning: finset.disj_Union_disj_Union -> Finset.disjUnionᵢ_disjUnionᵢ is a dubious translation: lean 3 declaration is - forall {α : Type.{u1}} {β : Type.{u2}} {γ : Type.{u3}} (s : Finset.{u1} α) (f : α -> (Finset.{u2} β)) (g : β -> (Finset.{u3} γ)) (h1 : Set.PairwiseDisjoint.{u2, u1} (Finset.{u2} β) α (Finset.partialOrder.{u2} β) (Finset.orderBot.{u2} β) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} α) (Set.{u1} α) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (Finset.Set.hasCoeT.{u1} α))) s) f) (h2 : Set.PairwiseDisjoint.{u3, u2} (Finset.{u3} γ) β (Finset.partialOrder.{u3} γ) (Finset.orderBot.{u3} γ) ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (Finset.disjUnionₓ.{u1, u2} α β s f h1)) g), Eq.{succ u3} (Finset.{u3} γ) (Finset.disjUnionₓ.{u2, u3} β γ (Finset.disjUnionₓ.{u1, u2} α β s f h1) g h2) (Finset.disjUnionₓ.{u1, u3} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) γ (Finset.attach.{u1} α s) (fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) (fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) (ha : Membership.Mem.{u1, u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.hasMem.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) a ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Finset.Set.hasCoeT.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) (Finset.attach.{u1} α s))) (b : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) (hb : Membership.Mem.{u1, u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.hasMem.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) b ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Finset.Set.hasCoeT.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) (Finset.attach.{u1} α s))) (hab : Ne.{succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) a b) => Iff.mpr (Disjoint.{u3} (Finset.{u3} γ) (Finset.partialOrder.{u3} γ) (Finset.orderBot.{u3} γ) ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) a) ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) b)) (forall {{a_1 : γ}}, (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a_1 ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) a)) -> (Not (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a_1 ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) b)))) (Finset.disjoint_left.{u3} γ ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) a) ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) b)) (fun (x : γ) (hxa : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) a)) (hxb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) b)) => Exists.dcases_on.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a_1))) (fun (_fresh.579.8860 : Exists.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a_1)))) => False) (Iff.mp (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc)))))) (Exists.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a_1)))) (Finset.mem_disjUnionᵢ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g x (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) hxa) (fun (xa : β) (h : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa))) => Exists.dcases_on.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa)) (fun (h : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa))) => False) h (fun (hfa : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (hga : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa)) => Exists.dcases_on.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a))) (fun (_fresh.579.8977 : Exists.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a)))) => False) (Iff.mp (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)) g (fun (b_1 : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b_1 ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)))) => h2 b_1 (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b_1 h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hc)))))) (Exists.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a)))) (Finset.mem_disjUnionᵢ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)) g x (fun (b_1 : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b_1 ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)))) => h2 b_1 (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b_1 h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hc))))) hxb) (fun (xb : β) (h : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb))) => Exists.dcases_on.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb)) (fun (h : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb))) => False) h (fun (hfb : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (hgb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb)) => Iff.mp (Disjoint.{u3} (Finset.{u3} γ) (Finset.partialOrder.{u3} γ) (Finset.orderBot.{u3} γ) (g xa) (g xb)) (forall {{a : γ}}, (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a (g xa)) -> (Not (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a (g xb)))) (Finset.disjoint_left.{u3} γ (g xa) (g xb)) (h2 xa (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f xa h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hfa))) xb (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f xb h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hfb))) (id.{0} (Ne.{succ u2} β xa xb) (fun (ᾰ : Eq.{succ u2} β xa xb) => Eq.ndrec.{0, succ u2} β xa (fun (xb : β) => (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) -> (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb)) -> False) (fun (hfb : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (hgb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa)) => Iff.mp (Disjoint.{u2} (Finset.{u2} β) (Finset.partialOrder.{u2} β) (Finset.orderBot.{u2} β) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (forall {{a_1 : β}}, (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) -> (Not (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))))) (Finset.disjoint_left.{u2} β (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (h1 ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) (Function.Injective.ne.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)))))) (Subtype.coe_injective.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) a b hab)) xa hfa hfb) xb ᾰ hfb hgb))) x hga hgb))))))) + forall {α : Type.{u1}} {β : Type.{u2}} {γ : Type.{u3}} (s : Finset.{u1} α) (f : α -> (Finset.{u2} β)) (g : β -> (Finset.{u3} γ)) (h1 : Set.PairwiseDisjoint.{u2, u1} (Finset.{u2} β) α (Finset.partialOrder.{u2} β) (Finset.orderBot.{u2} β) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} α) (Set.{u1} α) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (Finset.Set.hasCoeT.{u1} α))) s) f) (h2 : Set.PairwiseDisjoint.{u3, u2} (Finset.{u3} γ) β (Finset.partialOrder.{u3} γ) (Finset.orderBot.{u3} γ) ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (Finset.disjUnionₓ.{u1, u2} α β s f h1)) g), Eq.{succ u3} (Finset.{u3} γ) (Finset.disjUnionₓ.{u2, u3} β γ (Finset.disjUnionₓ.{u1, u2} α β s f h1) g h2) (Finset.disjUnionₓ.{u1, u3} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) γ (Finset.attach.{u1} α s) (fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) (fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) (ha : Membership.Mem.{u1, u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.hasMem.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) a ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Finset.Set.hasCoeT.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) (Finset.attach.{u1} α s))) (b : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) (hb : Membership.Mem.{u1, u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.hasMem.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) b ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Set.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))) (Finset.Set.hasCoeT.{u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) (Finset.attach.{u1} α s))) (hab : Ne.{succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) a b) => Iff.mpr (Disjoint.{u3} (Finset.{u3} γ) (Finset.partialOrder.{u3} γ) (Finset.orderBot.{u3} γ) ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) a) ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) b)) (forall {{a_1 : γ}}, (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a_1 ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) a)) -> (Not (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a_1 ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) b)))) (Finset.disjoint_left.{u3} γ ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) a) ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) b)) (fun (x : γ) (hxa : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) a)) (hxb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) => Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) b)) => Exists.dcases_on.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a_1))) (fun (_fresh.577.185166 : Exists.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a_1)))) => False) (Iff.mp (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc)))))) (Exists.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a_1)))) (Finset.mem_disjUnionᵢ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) g x (fun (b : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)))) => h2 b (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hc))))) hxa) (fun (xa : β) (h : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa))) => Exists.dcases_on.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa)) (fun (h : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa))) => False) h (fun (hfa : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (hga : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa)) => Exists.dcases_on.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a))) (fun (_fresh.577.185283 : Exists.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a)))) => False) (Iff.mp (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (Finset.disjUnionₓ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)) g (fun (b_1 : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b_1 ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)))) => h2 b_1 (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b_1 h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hc)))))) (Exists.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g a)))) (Finset.mem_disjUnionᵢ.{u2, u3} β γ (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)) g x (fun (b_1 : β) (hb : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) b_1 ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)))) (c : β) (hc : Membership.Mem.{u2, u2} β (Set.{u2} β) (Set.hasMem.{u2} β) c ((fun (a : Type.{u2}) (b : Type.{u2}) [self : HasLiftT.{succ u2, succ u2} a b] => self.0) (Finset.{u2} β) (Set.{u2} β) (HasLiftT.mk.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (CoeTCₓ.coe.{succ u2, succ u2} (Finset.{u2} β) (Set.{u2} β) (Finset.Set.hasCoeT.{u2} β))) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b)))) => h2 b_1 (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f b_1 h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) b_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hb))) c (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f c h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) c (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hc))))) hxb) (fun (xb : β) (h : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb))) => Exists.dcases_on.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb)) (fun (h : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) => Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb))) => False) h (fun (hfb : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (hgb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb)) => Iff.mp (Disjoint.{u3} (Finset.{u3} γ) (Finset.partialOrder.{u3} γ) (Finset.orderBot.{u3} γ) (g xa) (g xb)) (forall {{a : γ}}, (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a (g xa)) -> (Not (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a (g xb)))) (Finset.disjoint_left.{u3} γ (g xa) (g xb)) (h2 xa (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f xa h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) hfa))) xb (Iff.mpr (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (Finset.disjUnionₓ.{u1, u2} α β s f h1)) (Exists.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f a)))) (Finset.mem_disjUnionᵢ.{u1, u2} α β s f xb h1) (Exists.intro.{succ u1} α (fun (a : α) => Exists.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) a s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f a))) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Exists.intro.{0} (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) (fun (H : Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) s) => Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) hfb))) (id.{0} (Ne.{succ u2} β xa xb) (fun (ᾰ : Eq.{succ u2} β xa xb) => Eq.ndrec.{0, succ u2} β xa (fun (xb : β) => (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) -> (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xb)) -> False) (fun (hfb : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (hgb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (g xa)) => Iff.mp (Disjoint.{u2} (Finset.{u2} β) (Finset.partialOrder.{u2} β) (Finset.orderBot.{u2} β) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (forall {{a_1 : β}}, (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a))) -> (Not (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))))) (Finset.disjoint_left.{u2} β (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a)) (f ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b))) (h1 ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) a) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) a) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s))))) b) (Subtype.prop.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) b) (Function.Injective.ne.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (HasLiftT.mk.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (CoeTCₓ.coe.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeBase.{succ u1, succ u1} (Subtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) α (coeSubtype.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)))))) (Subtype.coe_injective.{succ u1} α (fun (x : α) => Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s)) a b hab)) xa hfa hfb) xb ᾰ hfb hgb))) x hga hgb))))))) but is expected to have type - forall {α : Type.{u3}} {β : Type.{u2}} {γ : Type.{u1}} (s : Finset.{u3} α) (f : α -> (Finset.{u2} β)) (g : β -> (Finset.{u1} γ)) (h1 : Set.PairwiseDisjoint.{u2, u3} (Finset.{u2} β) α (Finset.partialOrder.{u2} β) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u2} β) (Finset.toSet.{u3} α s) f) (h2 : Set.PairwiseDisjoint.{u1, u2} (Finset.{u1} γ) β (Finset.partialOrder.{u1} γ) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u1} γ) (Finset.toSet.{u2} β (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) g), Eq.{succ u1} (Finset.{u1} γ) (Finset.disjUnionᵢ.{u2, u1} β γ (Finset.disjUnionᵢ.{u3, u2} α β s f h1) g h2) (Finset.disjUnionᵢ.{u3, u1} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) γ (Finset.attach.{u3} α s) (fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) (fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (ha : Membership.mem.{u3, u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (Set.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s))) (Set.instMembershipSet.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s))) a (Finset.toSet.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (Finset.attach.{u3} α s))) (b : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (hb : Membership.mem.{u3, u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (Set.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s))) (Set.instMembershipSet.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s))) b (Finset.toSet.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (Finset.attach.{u3} α s))) (hab : Ne.{succ u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) a b) => Iff.mpr (Disjoint.{u1} (Finset.{u1} γ) (Finset.partialOrder.{u1} γ) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u1} γ) ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) a) ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) b)) (forall {{a_1 : γ}}, (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a_1 ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) a)) -> (Not (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a_1 ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) b)))) (Finset.disjoint_left.{u1} γ ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) a) ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) b)) (fun (x : γ) (hxa : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) a)) (hxb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) b)) => Exists.casesOn.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a_1))) (fun (_fresh.579.8860 : Exists.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a_1)))) => False) (Iff.mp (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc)))))) (Exists.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a_1)))) (Finset.mem_disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g x (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) hxa) (fun (xa : β) (h : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa))) => And.casesOn.{0} (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa)) (fun (h : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa))) => False) h (fun (hfa : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (hga : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa)) => Exists.casesOn.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a))) (fun (_fresh.579.8977 : Exists.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a)))) => False) (Iff.mp (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)) g (fun (b_1 : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b_1 (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)))) => h2 b_1 (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b_1 h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hc)))))) (Exists.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a)))) (Finset.mem_disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)) g x (fun (b_1 : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b_1 (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)))) => h2 b_1 (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b_1 h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hc))))) hxb) (fun (xb : β) (h : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb))) => And.casesOn.{0} (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb)) (fun (h : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb))) => False) h (fun (hfb : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (hgb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb)) => Iff.mp (Disjoint.{u1} (Finset.{u1} γ) (Finset.partialOrder.{u1} γ) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u1} γ) (g xa) (g xb)) (forall {{a : γ}}, (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a (g xa)) -> (Not (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a (g xb)))) (Finset.disjoint_left.{u1} γ (g xa) (g xb)) (h2 xa (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f xa h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hfa))) xb (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f xb h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hfb))) (fun (a._@.Init.Prelude.139.Mathlib.Data.Finset.Basic._hyg.32954 : Eq.{succ u2} β xa xb) => Eq.ndrec.{0, succ u2} β xa (fun (xb : β) => (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) -> (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb)) -> False) (fun (hfb : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (hgb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa)) => Iff.mp (Disjoint.{u2} (Finset.{u2} β) (Finset.partialOrder.{u2} β) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u2} β) (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (forall {{a_1 : β}}, (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) -> (Not (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))))) (Finset.disjoint_left.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (h1 (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (Function.Injective.ne.{succ u3, succ u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) α (fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (Subtype.coe_injective.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) a b hab)) xa hfa hfb) xb a._@.Init.Prelude.139.Mathlib.Data.Finset.Basic._hyg.32954 hfb hgb)) x hga hgb))))))) + forall {α : Type.{u3}} {β : Type.{u2}} {γ : Type.{u1}} (s : Finset.{u3} α) (f : α -> (Finset.{u2} β)) (g : β -> (Finset.{u1} γ)) (h1 : Set.PairwiseDisjoint.{u2, u3} (Finset.{u2} β) α (Finset.partialOrder.{u2} β) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u2} β) (Finset.toSet.{u3} α s) f) (h2 : Set.PairwiseDisjoint.{u1, u2} (Finset.{u1} γ) β (Finset.partialOrder.{u1} γ) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u1} γ) (Finset.toSet.{u2} β (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) g), Eq.{succ u1} (Finset.{u1} γ) (Finset.disjUnionᵢ.{u2, u1} β γ (Finset.disjUnionᵢ.{u3, u2} α β s f h1) g h2) (Finset.disjUnionᵢ.{u3, u1} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) γ (Finset.attach.{u3} α s) (fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) (fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (ha : Membership.mem.{u3, u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (Set.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s))) (Set.instMembershipSet.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s))) a (Finset.toSet.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (Finset.attach.{u3} α s))) (b : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (hb : Membership.mem.{u3, u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (Set.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s))) (Set.instMembershipSet.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s))) b (Finset.toSet.{u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) (Finset.attach.{u3} α s))) (hab : Ne.{succ u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) a b) => Iff.mpr (Disjoint.{u1} (Finset.{u1} γ) (Finset.partialOrder.{u1} γ) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u1} γ) ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) a) ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) b)) (forall {{a_1 : γ}}, (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a_1 ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) a)) -> (Not (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a_1 ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) b)))) (Finset.disjoint_left.{u1} γ ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) a) ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) b)) (fun (x : γ) (hxa : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) a)) (hxb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) b)) => Exists.casesOn.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a_1))) (fun (_fresh.577.185166 : Exists.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a_1)))) => False) (Iff.mp (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc)))))) (Exists.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a_1)))) (Finset.mem_disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) g x (fun (b : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)))) => h2 b (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hc))))) hxa) (fun (xa : β) (h : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa))) => And.casesOn.{0} (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa)) (fun (h : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa))) => False) h (fun (hfa : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (hga : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa)) => Exists.casesOn.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a))) (fun (_fresh.577.185283 : Exists.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a)))) => False) (Iff.mp (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (Finset.disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)) g (fun (b_1 : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b_1 (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)))) => h2 b_1 (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b_1 h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hc)))))) (Exists.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g a)))) (Finset.mem_disjUnionᵢ.{u2, u1} β γ (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)) g x (fun (b_1 : β) (hb : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) b_1 (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)))) (c : β) (hc : Membership.mem.{u2, u2} β (Set.{u2} β) (Set.instMembershipSet.{u2} β) c (Finset.toSet.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b)))) => h2 b_1 (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f b_1 h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) b_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hb))) c (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f c h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) c (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hc))))) hxb) (fun (xb : β) (h : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb))) => And.casesOn.{0} (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb)) (fun (h : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb))) => False) h (fun (hfb : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (hgb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb)) => Iff.mp (Disjoint.{u1} (Finset.{u1} γ) (Finset.partialOrder.{u1} γ) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u1} γ) (g xa) (g xb)) (forall {{a : γ}}, (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a (g xa)) -> (Not (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a (g xb)))) (Finset.disjoint_left.{u1} γ (g xa) (g xb)) (h2 xa (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f xa h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) hfa))) xb (Iff.mpr (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (Finset.disjUnionᵢ.{u3, u2} α β s f h1)) (Exists.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f a)))) (Finset.mem_disjUnionᵢ.{u3, u2} α β s f xb h1) (Exists.intro.{succ u3} α (fun (a : α) => And (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) a s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f a))) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (And.intro (Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) s) (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) hfb))) (fun (a._@.Init.Prelude.139.Mathlib.Data.Finset.Basic._hyg.32954 : Eq.{succ u2} β xa xb) => Eq.ndrec.{0, succ u2} β xa (fun (xb : β) => (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) -> (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xb)) -> False) (fun (hfb : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (hgb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (g xa)) => Iff.mp (Disjoint.{u2} (Finset.{u2} β) (Finset.partialOrder.{u2} β) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u2} β) (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (forall {{a_1 : β}}, (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a))) -> (Not (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))))) (Finset.disjoint_left.{u2} β (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a)) (f (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b))) (h1 (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (Subtype.prop.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) b) (Function.Injective.ne.{succ u3, succ u3} (Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) α (fun (a : Subtype.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) => Subtype.val.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s) a) (Subtype.coe_injective.{succ u3} α (fun (x : α) => Membership.mem.{u3, u3} α (Finset.{u3} α) (Finset.instMembershipFinset.{u3} α) x s)) a b hab)) xa hfa hfb) xb a._@.Init.Prelude.139.Mathlib.Data.Finset.Basic._hyg.32954 hfb hgb)) x hga hgb))))))) Case conversion may be inaccurate. Consider using '#align finset.disj_Union_disj_Union Finset.disjUnionᵢ_disjUnionᵢₓ'. -/ theorem disjUnionᵢ_disjUnionᵢ (s : Finset α) (f : α → Finset β) (g : β → Finset γ) (h1 h2) : (s.disjUnionₓ f h1).disjUnionₓ g h2 = diff --git a/Mathbin/Data/Finset/Image.lean b/Mathbin/Data/Finset/Image.lean index 2f43cc9ecc..f59f238f51 100644 --- a/Mathbin/Data/Finset/Image.lean +++ b/Mathbin/Data/Finset/Image.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Minchao Wu, Mario Carneiro ! This file was ported from Lean 3 source module data.finset.image -! leanprover-community/mathlib commit e04043d6bf7264a3c84bc69711dc354958ca4516 +! leanprover-community/mathlib commit b685f506164f8d17a6404048bc4d696739c5d976 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -487,9 +487,9 @@ theorem map_disjUnionᵢ {f : α ↪ β} {s : Finset α} {t : β → Finset γ} /- warning: finset.disj_Union_map -> Finset.disjUnionᵢ_map is a dubious translation: lean 3 declaration is - forall {α : Type.{u1}} {β : Type.{u2}} {γ : Type.{u3}} {s : Finset.{u1} α} {t : α -> (Finset.{u2} β)} {f : Function.Embedding.{succ u2, succ u3} β γ} {h : Set.PairwiseDisjoint.{u2, u1} (Finset.{u2} β) α (Finset.partialOrder.{u2} β) (Finset.orderBot.{u2} β) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} α) (Set.{u1} α) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (Finset.Set.hasCoeT.{u1} α))) s) t}, Eq.{succ u3} (Finset.{u3} γ) (Finset.map.{u2, u3} β γ f (Finset.disjUnionₓ.{u1, u2} α β s t h)) (Finset.disjUnionₓ.{u1, u3} α γ s (fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) (fun (a : α) (ha : Membership.Mem.{u1, u1} α (Set.{u1} α) (Set.hasMem.{u1} α) a ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} α) (Set.{u1} α) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (Finset.Set.hasCoeT.{u1} α))) s)) (b : α) (hb : Membership.Mem.{u1, u1} α (Set.{u1} α) (Set.hasMem.{u1} α) b ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} α) (Set.{u1} α) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (Finset.Set.hasCoeT.{u1} α))) s)) (hab : Ne.{succ u1} α a b) => Iff.mpr (Disjoint.{u3} (Finset.{u3} γ) (Finset.partialOrder.{u3} γ) (Finset.orderBot.{u3} γ) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) (forall {{a_1 : γ}}, (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a_1 ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) -> (Not (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a_1 ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)))) (Finset.disjoint_left.{u3} γ ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) (fun (x : γ) (hxa : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) (hxb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) => Exists.dcases_on.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a_1) x)) (fun (_fresh.582.24927 : Exists.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a_1) x))) => False) (Iff.mp (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (Finset.map.{u2, u3} β γ f (t a))) (Exists.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a_1) x))) (Finset.mem_map.{u2, u3} β γ f (t a) x) hxa) (fun (xa : β) (h_1 : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) x)) => Exists.dcases_on.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) x) (fun (h_1 : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) x)) => False) h_1 (fun (hfa : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) (h_1_h : Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) x) => Eq.ndrec.{0, succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) (fun (x : γ) => (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) -> (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) -> False) (fun (hxa : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) (hxb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) => Exists.dcases_on.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa))) (fun (_fresh.582.25013 : Exists.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)))) => False) (Iff.mp (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) (Finset.map.{u2, u3} β γ f (t b))) (Exists.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)))) (Finset.mem_map.{u2, u3} β γ f (t b) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)) hxb) (fun (xb : β) (h_1 : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa))) => Exists.dcases_on.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)) (fun (h_1 : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa))) => False) h_1 (fun (hfb : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) (hfab : Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)) => Eq.ndrec.{0, succ u2} β xb (fun (xa : β) => (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) -> (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) -> (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) -> (Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)) -> False) (fun (hfa : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t a)) (hxa : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) (hxb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) (hfab : Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb)) => Iff.mp (Disjoint.{u2} (Finset.{u2} β) (Finset.partialOrder.{u2} β) (Finset.orderBot.{u2} β) (t a) (t b)) (forall {{a_1 : β}}, (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) -> (Not (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t b)))) (Finset.disjoint_left.{u2} β (t a) (t b)) (h a ha b hb hab) xb hfa hfb) xa (Function.Embedding.injective.{succ u2, succ u3} β γ f xb xa hfab) hfa hxa hxb hfab))) x h_1_h hxa hxb))))) + forall {α : Type.{u1}} {β : Type.{u2}} {γ : Type.{u3}} {s : Finset.{u1} α} {t : α -> (Finset.{u2} β)} {f : Function.Embedding.{succ u2, succ u3} β γ} {h : Set.PairwiseDisjoint.{u2, u1} (Finset.{u2} β) α (Finset.partialOrder.{u2} β) (Finset.orderBot.{u2} β) ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} α) (Set.{u1} α) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (Finset.Set.hasCoeT.{u1} α))) s) t}, Eq.{succ u3} (Finset.{u3} γ) (Finset.map.{u2, u3} β γ f (Finset.disjUnionₓ.{u1, u2} α β s t h)) (Finset.disjUnionₓ.{u1, u3} α γ s (fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) (fun (a : α) (ha : Membership.Mem.{u1, u1} α (Set.{u1} α) (Set.hasMem.{u1} α) a ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} α) (Set.{u1} α) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (Finset.Set.hasCoeT.{u1} α))) s)) (b : α) (hb : Membership.Mem.{u1, u1} α (Set.{u1} α) (Set.hasMem.{u1} α) b ((fun (a : Type.{u1}) (b : Type.{u1}) [self : HasLiftT.{succ u1, succ u1} a b] => self.0) (Finset.{u1} α) (Set.{u1} α) (HasLiftT.mk.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (CoeTCₓ.coe.{succ u1, succ u1} (Finset.{u1} α) (Set.{u1} α) (Finset.Set.hasCoeT.{u1} α))) s)) (hab : Ne.{succ u1} α a b) => Iff.mpr (Disjoint.{u3} (Finset.{u3} γ) (Finset.partialOrder.{u3} γ) (Finset.orderBot.{u3} γ) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) (forall {{a_1 : γ}}, (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a_1 ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) -> (Not (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) a_1 ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)))) (Finset.disjoint_left.{u3} γ ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) (fun (x : γ) (hxa : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) (hxb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) => Exists.dcases_on.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a_1) x)) (fun (_fresh.607.3104 : Exists.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a_1) x))) => False) (Iff.mp (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x (Finset.map.{u2, u3} β γ f (t a))) (Exists.{succ u2} β (fun (a_1 : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a_1) x))) (Finset.mem_map.{u2, u3} β γ f (t a) x) hxa) (fun (xa : β) (h_1 : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) x)) => Exists.dcases_on.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) x) (fun (h_1 : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) x)) => False) h_1 (fun (hfa : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) (h_1_h : Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) x) => Eq.ndrec.{0, succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) (fun (x : γ) => (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) -> (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) x ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) -> False) (fun (hxa : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) (hxb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) => Exists.dcases_on.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa))) (fun (_fresh.607.3190 : Exists.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)))) => False) (Iff.mp (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) (Finset.map.{u2, u3} β γ f (t b))) (Exists.{succ u2} β (fun (a : β) => Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f a) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)))) (Finset.mem_map.{u2, u3} β γ f (t b) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)) hxb) (fun (xb : β) (h_1 : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa))) => Exists.dcases_on.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)) (fun (h_1 : Exists.{0} (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) (fun (H : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) => Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa))) => False) h_1 (fun (hfb : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t b)) (hfab : Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)) => Eq.ndrec.{0, succ u2} β xb (fun (xa : β) => (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xa (t a)) -> (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) -> (Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) -> (Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xa)) -> False) (fun (hfa : Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) xb (t a)) (hxa : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) a)) (hxb : Membership.Mem.{u3, u3} γ (Finset.{u3} γ) (Finset.hasMem.{u3} γ) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) ((fun (a : α) => Finset.map.{u2, u3} β γ f (t a)) b)) (hfab : Eq.{succ u3} γ (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb) (coeFn.{max 1 (succ u2) (succ u3), max (succ u2) (succ u3)} (Function.Embedding.{succ u2, succ u3} β γ) (fun (_x : Function.Embedding.{succ u2, succ u3} β γ) => β -> γ) (Function.Embedding.hasCoeToFun.{succ u2, succ u3} β γ) f xb)) => Iff.mp (Disjoint.{u2} (Finset.{u2} β) (Finset.partialOrder.{u2} β) (Finset.orderBot.{u2} β) (t a) (t b)) (forall {{a_1 : β}}, (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t a)) -> (Not (Membership.Mem.{u2, u2} β (Finset.{u2} β) (Finset.hasMem.{u2} β) a_1 (t b)))) (Finset.disjoint_left.{u2} β (t a) (t b)) (h a ha b hb hab) xb hfa hfb) xa (Function.Embedding.injective.{succ u2, succ u3} β γ f xb xa hfab) hfa hxa hxb hfab))) x h_1_h hxa hxb))))) but is expected to have type - forall {α : Type.{u3}} {β : Type.{u2}} {γ : Type.{u1}} {s : Finset.{u3} α} {t : α -> (Finset.{u2} β)} {f : Function.Embedding.{succ u2, succ u1} β γ} {h : Set.PairwiseDisjoint.{u2, u3} (Finset.{u2} β) α (Finset.partialOrder.{u2} β) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u2} β) (Finset.toSet.{u3} α s) t}, Eq.{succ u1} (Finset.{u1} γ) (Finset.map.{u2, u1} β γ f (Finset.disjUnionᵢ.{u3, u2} α β s t h)) (Finset.disjUnionᵢ.{u3, u1} α γ s (fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) (fun (a : α) (ha : Membership.mem.{u3, u3} α (Set.{u3} α) (Set.instMembershipSet.{u3} α) a (Finset.toSet.{u3} α s)) (b : α) (hb : Membership.mem.{u3, u3} α (Set.{u3} α) (Set.instMembershipSet.{u3} α) b (Finset.toSet.{u3} α s)) (hab : Ne.{succ u3} α a b) => Iff.mpr (Disjoint.{u1} (Finset.{u1} γ) (Finset.partialOrder.{u1} γ) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u1} γ) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) (forall {{a_1 : γ}}, (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a_1 ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) -> (Not (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a_1 ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)))) (Finset.disjoint_left.{u1} γ ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) (fun (x : γ) (hxa : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) (hxb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) => Exists.casesOn.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a_1) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a_1) x)) (fun (_fresh.582.24927 : Exists.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a_1) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a_1) x))) => False) (Iff.mp (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (Finset.map.{u2, u1} β γ f (t a))) (Exists.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a_1) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a_1) x))) (Finset.mem_map.{u2, u1} β γ f (t a) x) hxa) (fun (xa : β) (h_1 : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) x)) => And.casesOn.{0} (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) x) (fun (h_1 : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) x)) => False) h_1 (fun (hfa : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) (h_1_h : Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) x) => Eq.ndrec.{0, succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) (fun (x : γ) => (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) -> (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) -> False) (fun (hxa : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) (hxb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) => Exists.casesOn.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa))) (fun (_fresh.582.25013 : Exists.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)))) => False) (Iff.mp (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) (Finset.map.{u2, u1} β γ f (t b))) (Exists.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)))) (Finset.mem_map.{u2, u1} β γ f (t b) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)) hxb) (fun (xb : β) (h_1 : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa))) => And.casesOn.{0} (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)) (fun (h_1 : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa))) => False) h_1 (fun (hfb : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t b)) (hfab : Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)) => Eq.ndrec.{0, succ u2} β xb (fun (xa : β) => (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) -> (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) -> (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) -> (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)) -> False) (fun (hfa : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t a)) (hxa : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) (hxb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) (hfab : Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb)) => Iff.mp (Disjoint.{u2} (Finset.{u2} β) (Finset.partialOrder.{u2} β) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u2} β) (t a) (t b)) (forall {{a_1 : β}}, (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t a)) -> (Not (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t b)))) (Finset.disjoint_left.{u2} β (t a) (t b)) (h a ha b hb hab) xb hfa hfb) xa (Function.Embedding.injective.{succ u1, succ u2} β γ f xb xa hfab) hfa hxa hxb hfab))) x h_1_h hxa hxb))))) + forall {α : Type.{u3}} {β : Type.{u2}} {γ : Type.{u1}} {s : Finset.{u3} α} {t : α -> (Finset.{u2} β)} {f : Function.Embedding.{succ u2, succ u1} β γ} {h : Set.PairwiseDisjoint.{u2, u3} (Finset.{u2} β) α (Finset.partialOrder.{u2} β) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u2} β) (Finset.toSet.{u3} α s) t}, Eq.{succ u1} (Finset.{u1} γ) (Finset.map.{u2, u1} β γ f (Finset.disjUnionᵢ.{u3, u2} α β s t h)) (Finset.disjUnionᵢ.{u3, u1} α γ s (fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) (fun (a : α) (ha : Membership.mem.{u3, u3} α (Set.{u3} α) (Set.instMembershipSet.{u3} α) a (Finset.toSet.{u3} α s)) (b : α) (hb : Membership.mem.{u3, u3} α (Set.{u3} α) (Set.instMembershipSet.{u3} α) b (Finset.toSet.{u3} α s)) (hab : Ne.{succ u3} α a b) => Iff.mpr (Disjoint.{u1} (Finset.{u1} γ) (Finset.partialOrder.{u1} γ) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u1} γ) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) (forall {{a_1 : γ}}, (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a_1 ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) -> (Not (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) a_1 ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)))) (Finset.disjoint_left.{u1} γ ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) (fun (x : γ) (hxa : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) (hxb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) => Exists.casesOn.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a_1) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a_1) x)) (fun (_fresh.607.3104 : Exists.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a_1) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a_1) x))) => False) (Iff.mp (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x (Finset.map.{u2, u1} β γ f (t a))) (Exists.{succ u2} β (fun (a_1 : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a_1) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a_1) x))) (Finset.mem_map.{u2, u1} β γ f (t a) x) hxa) (fun (xa : β) (h_1 : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) x)) => And.casesOn.{0} (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) x) (fun (h_1 : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) x)) => False) h_1 (fun (hfa : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) (h_1_h : Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) x) => Eq.ndrec.{0, succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xa) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) (fun (x : γ) => (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) -> (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) x ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) -> False) (fun (hxa : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) (hxb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) => Exists.casesOn.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa))) (fun (_fresh.607.3190 : Exists.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)))) => False) (Iff.mp (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) (Finset.map.{u2, u1} β γ f (t b))) (Exists.{succ u2} β (fun (a : β) => And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f a) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)))) (Finset.mem_map.{u2, u1} β γ f (t b) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)) hxb) (fun (xb : β) (h_1 : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa))) => And.casesOn.{0} (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)) (fun (h_1 : And (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t b)) (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (a : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) a) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa))) => False) h_1 (fun (hfb : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t b)) (hfab : Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)) => Eq.ndrec.{0, succ u2} β xb (fun (xa : β) => (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xa (t a)) -> (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) -> (Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) -> (Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xa)) -> False) (fun (hfa : Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) xb (t a)) (hxa : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) a)) (hxb : Membership.mem.{u1, u1} γ (Finset.{u1} γ) (Finset.instMembershipFinset.{u1} γ) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) ((fun (a : α) => Finset.map.{u2, u1} β γ f (t a)) b)) (hfab : Eq.{succ u1} ((fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb) (FunLike.coe.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β (fun (_x : β) => (fun (x._@.Mathlib.Data.FunLike.Embedding._hyg.19 : β) => γ) _x) (EmbeddingLike.toFunLike.{max (succ u2) (succ u1), succ u2, succ u1} (Function.Embedding.{succ u2, succ u1} β γ) β γ (Function.instEmbeddingLikeEmbedding.{succ u2, succ u1} β γ)) f xb)) => Iff.mp (Disjoint.{u2} (Finset.{u2} β) (Finset.partialOrder.{u2} β) (Finset.instOrderBotFinsetToLEToPreorderPartialOrder.{u2} β) (t a) (t b)) (forall {{a_1 : β}}, (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t a)) -> (Not (Membership.mem.{u2, u2} β (Finset.{u2} β) (Finset.instMembershipFinset.{u2} β) a_1 (t b)))) (Finset.disjoint_left.{u2} β (t a) (t b)) (h a ha b hb hab) xb hfa hfb) xa (Function.Embedding.injective.{succ u1, succ u2} β γ f xb xa hfab) hfa hxa hxb hfab))) x h_1_h hxa hxb))))) Case conversion may be inaccurate. Consider using '#align finset.disj_Union_map Finset.disjUnionᵢ_mapₓ'. -/ theorem disjUnionᵢ_map {s : Finset α} {t : α → Finset β} {f : β ↪ γ} {h} : (s.disjUnionₓ t h).map f = @@ -840,11 +840,9 @@ but is expected to have type Case conversion may be inaccurate. Consider using '#align finset.image_inter_of_inj_on Finset.image_inter_of_injOnₓ'. -/ theorem image_inter_of_injOn [DecidableEq α] {f : α → β} (s t : Finset α) (hf : Set.InjOn f (s ∪ t)) : (s ∩ t).image f = s.image f ∩ t.image f := - (image_inter_subset _ _ _).antisymm fun x => - by - simp only [mem_inter, mem_image] - rintro ⟨⟨a, ha, rfl⟩, b, hb, h⟩ - exact ⟨a, ⟨ha, by rwa [← hf (Or.inr hb) (Or.inl ha) h]⟩, rfl⟩ + coe_injective <| by + push_cast + exact Set.image_inter_on fun a ha b hb => hf (Or.inr ha) <| Or.inl hb #align finset.image_inter_of_inj_on Finset.image_inter_of_injOn /- warning: finset.image_inter -> Finset.image_inter is a dubious translation: @@ -915,6 +913,20 @@ theorem image_eq_empty : s.image f = ∅ ↔ s = ∅ := #align finset.image_eq_empty Finset.image_eq_empty -/ +theorem image_sdiff [DecidableEq α] {f : α → β} (s t : Finset α) (hf : Injective f) : + (s \ t).image f = s.image f \ t.image f := + coe_injective <| by + push_cast + exact Set.image_diff hf _ _ +#align finset.image_sdiff Finset.image_sdiff + +theorem image_symmDiff [DecidableEq α] {f : α → β} (s t : Finset α) (hf : Injective f) : + (s ∆ t).image f = s.image f ∆ t.image f := + coe_injective <| by + push_cast + exact Set.image_symm_diff hf _ _ +#align finset.image_symm_diff Finset.image_symmDiff + /- warning: disjoint.of_image_finset -> Disjoint.of_image_finset is a dubious translation: lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : DecidableEq.{succ u2} β] {s : Finset.{u1} α} {t : Finset.{u1} α} {f : α -> β}, (Disjoint.{u2} (Finset.{u2} β) (Finset.partialOrder.{u2} β) (Finset.orderBot.{u2} β) (Finset.image.{u1, u2} α β (fun (a : β) (b : β) => _inst_1 a b) f s) (Finset.image.{u1, u2} α β (fun (a : β) (b : β) => _inst_1 a b) f t)) -> (Disjoint.{u1} (Finset.{u1} α) (Finset.partialOrder.{u1} α) (Finset.orderBot.{u1} α) s t) diff --git a/Mathbin/Data/Finset/Pointwise.lean b/Mathbin/Data/Finset/Pointwise.lean index 43a34e5797..5968caf6f4 100644 --- a/Mathbin/Data/Finset/Pointwise.lean +++ b/Mathbin/Data/Finset/Pointwise.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Yaël Dillies ! This file was ported from Lean 3 source module data.finset.pointwise -! leanprover-community/mathlib commit 517cc149e0b515d2893baa376226ed10feb319c7 +! leanprover-community/mathlib commit b685f506164f8d17a6404048bc4d696739c5d976 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -2519,10 +2519,12 @@ instance isScalarTower'' [SMul α β] [SMul α γ] [SMul β γ] [IsScalarTower -/ #print Finset.isCentralScalar /- +@[to_additive] instance isCentralScalar [SMul α β] [SMul αᵐᵒᵖ β] [IsCentralScalar α β] : IsCentralScalar α (Finset β) := ⟨fun a s => coe_injective <| by simp only [coe_smul_finset, coe_smul, op_smul_eq_smul]⟩ #align finset.is_central_scalar Finset.isCentralScalar +#align finset.is_central_vadd Finset.is_central_vadd -/ #print Finset.mulAction /- @@ -2797,6 +2799,38 @@ theorem subset_smul_finset_iff : s ⊆ a • t ↔ a⁻¹ • s ⊆ t := #align finset.subset_smul_finset_iff Finset.subset_smul_finset_iff #align finset.subset_vadd_finset_iff Finset.subset_vadd_finset_iff +@[to_additive] +theorem smul_finset_inter : a • (s ∩ t) = a • s ∩ a • t := + image_inter _ _ <| MulAction.injective a +#align finset.smul_finset_inter Finset.smul_finset_inter +#align finset.vadd_finset_inter Finset.vadd_finset_inter + +@[to_additive] +theorem smul_finset_sdiff : a • (s \ t) = a • s \ a • t := + image_sdiff _ _ <| MulAction.injective a +#align finset.smul_finset_sdiff Finset.smul_finset_sdiff +#align finset.vadd_finset_sdiff Finset.vadd_finset_sdiff + +@[to_additive] +theorem smul_finset_symmDiff : a • s ∆ t = (a • s) ∆ (a • t) := + image_symmDiff _ _ <| MulAction.injective a +#align finset.smul_finset_symm_diff Finset.smul_finset_symmDiff +#align finset.vadd_finset_symm_diff Finset.vadd_finset_symmDiff + +@[simp, to_additive] +theorem smul_finset_univ [Fintype β] : a • (univ : Finset β) = univ := + image_univ_of_surjective <| MulAction.surjective a +#align finset.smul_finset_univ Finset.smul_finset_univ +#align finset.vadd_finset_univ Finset.vadd_finset_univ + +@[simp, to_additive] +theorem smul_univ [Fintype β] {s : Finset α} (hs : s.Nonempty) : s • (univ : Finset β) = univ := + coe_injective <| by + push_cast + exact Set.smul_univ hs +#align finset.smul_univ Finset.smul_univ +#align finset.vadd_univ Finset.vadd_univ + #print Finset.card_smul_finset /- @[simp, to_additive] theorem card_smul_finset (a : α) (s : Finset β) : (a • s).card = s.card := @@ -2873,6 +2907,18 @@ theorem subset_smul_finset_iff₀ (ha : a ≠ 0) : s ⊆ a • t ↔ a⁻¹ • show _ ⊆ Units.mk0 a ha • _ ↔ _ from subset_smul_finset_iff #align finset.subset_smul_finset_iff₀ Finset.subset_smul_finset_iff₀ +theorem smul_finset_inter₀ (ha : a ≠ 0) : a • (s ∩ t) = a • s ∩ a • t := + image_inter _ _ <| MulAction.injective₀ ha +#align finset.smul_finset_inter₀ Finset.smul_finset_inter₀ + +theorem smul_finset_sdiff₀ (ha : a ≠ 0) : a • (s \ t) = a • s \ a • t := + image_sdiff _ _ <| MulAction.injective₀ ha +#align finset.smul_finset_sdiff₀ Finset.smul_finset_sdiff₀ + +theorem smul_finset_symm_diff₀ (ha : a ≠ 0) : a • s ∆ t = (a • s) ∆ (a • t) := + image_symmDiff _ _ <| MulAction.injective₀ ha +#align finset.smul_finset_symm_diff₀ Finset.smul_finset_symm_diff₀ + /- warning: finset.smul_univ₀ -> Finset.smul_univ₀ is a dubious translation: lean 3 declaration is forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : DecidableEq.{succ u2} β] [_inst_2 : GroupWithZero.{u1} α] [_inst_3 : MulAction.{u1, u2} α β (MonoidWithZero.toMonoid.{u1} α (GroupWithZero.toMonoidWithZero.{u1} α _inst_2))] [_inst_4 : Fintype.{u2} β] {s : Finset.{u1} α}, (Not (HasSubset.Subset.{u1} (Finset.{u1} α) (Finset.hasSubset.{u1} α) s (OfNat.ofNat.{u1} (Finset.{u1} α) 0 (OfNat.mk.{u1} (Finset.{u1} α) 0 (Zero.zero.{u1} (Finset.{u1} α) (Finset.zero.{u1} α (MulZeroClass.toHasZero.{u1} α (MulZeroOneClass.toMulZeroClass.{u1} α (MonoidWithZero.toMulZeroOneClass.{u1} α (GroupWithZero.toMonoidWithZero.{u1} α _inst_2)))))))))) -> (Eq.{succ u2} (Finset.{u2} β) (SMul.smul.{u1, u2} (Finset.{u1} α) (Finset.{u2} β) (Finset.smul.{u1, u2} α β (fun (a : β) (b : β) => _inst_1 a b) (MulAction.toHasSmul.{u1, u2} α β (MonoidWithZero.toMonoid.{u1} α (GroupWithZero.toMonoidWithZero.{u1} α _inst_2)) _inst_3)) s (Finset.univ.{u2} β _inst_4)) (Finset.univ.{u2} β _inst_4)) diff --git a/Mathbin/Data/Seq/Parallel.lean b/Mathbin/Data/Seq/Parallel.lean index e5d7a8ab44..a1e59f5d18 100644 --- a/Mathbin/Data/Seq/Parallel.lean +++ b/Mathbin/Data/Seq/Parallel.lean @@ -11,7 +11,7 @@ terminates_parallel and exists_of_mem_parallel. honor sequence equivalence (irrelevance of computation time).) ! This file was ported from Lean 3 source module data.seq.parallel -! leanprover-community/mathlib commit 77615d00fbf05cd7e55ba84176070a18bc097772 +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -21,8 +21,8 @@ universe u v namespace Computation -open Wseq - +/- ./././Mathport/Syntax/Translate/Command.lean:224:11: unsupported: unusual advanced open style -/ +/- ./././Mathport/Syntax/Translate/Command.lean:224:11: unsupported: unusual advanced open style -/ variable {α : Type u} {β : Type v} /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @@ -42,8 +42,8 @@ def Parallel.aux1 : | (l, S) => rmap (fun l' => - match SeqCat.destruct S with - | none => (l', nil) + match Seq.destruct S with + | none => (l', Seq.nil) | some (none, S') => (l', S') | some (some c, S') => (c::l', S')) (Parallel.aux2 l) @@ -119,7 +119,7 @@ theorem TerminatesParallel.aux : rw [H2] apply @Computation.think_terminates _ _ _ have := H1 _ h - rcases SeqCat.destruct S with (_ | ⟨_ | c, S'⟩) <;> simp [parallel.aux1] <;> apply IH <;> + rcases seq.destruct S with (_ | ⟨_ | c, S'⟩) <;> simp [parallel.aux1] <;> apply IH <;> simp [this] #align computation.terminates_parallel.aux Computation.TerminatesParallel.aux @@ -128,17 +128,16 @@ theorem terminates_parallel {S : Wseq (Computation α)} {c} (h : c ∈ S) [T : T by suffices ∀ (n) (l : List (Computation α)) (S c), - c ∈ l ∨ some (some c) = SeqCat.nth S n → - Terminates c → Terminates (corec Parallel.aux1 (l, S)) + c ∈ l ∨ some (some c) = Seq.nth S n → Terminates c → Terminates (corec Parallel.aux1 (l, S)) from let ⟨n, h⟩ := h this n [] S c (Or.inr h) T intro n; induction' n with n IH <;> intro l S c o T · cases' o with a a · exact terminates_parallel.aux a T - have H : SeqCat.destruct S = some (some c, _) := + have H : seq.destruct S = some (some c, _) := by - unfold SeqCat.destruct Functor.map + unfold seq.destruct Functor.map rw [← a] simp induction' h : parallel.aux2 l with a l' <;> have C : corec parallel.aux1 (l, S) = _ @@ -180,19 +179,18 @@ theorem terminates_parallel {S : Wseq (Computation α)} {c} (h : c ∈ S) [T : T rw [a] cases' S with f al rfl - induction' e : SeqCat.nth S 0 with o - · have D : SeqCat.destruct S = none := - by - dsimp [SeqCat.destruct] + induction' e : seq.nth S 0 with o + · have D : seq.destruct S = none := by + dsimp [seq.destruct] rw [e] rfl rw [D] simp [parallel.aux1] have TT := TT l' - rwa [SeqCat.destruct_eq_nil D, SeqCat.tail_nil] at TT - · have D : SeqCat.destruct S = some (o, S.tail) := + rwa [seq.destruct_eq_nil D, seq.tail_nil] at TT + · have D : seq.destruct S = some (o, S.tail) := by - dsimp [SeqCat.destruct] + dsimp [seq.destruct] rw [e] rfl rw [D] @@ -252,11 +250,11 @@ theorem exists_of_mem_parallel {S : Wseq (Computation α)} {a} (h : a ∈ parall · rw [h'] at this rcases this with ⟨c, cl, ac⟩ exact ⟨c, Or.inl cl, ac⟩ - · induction' e : SeqCat.destruct S with a <;> rw [e] at h' + · induction' e : seq.destruct S with a <;> rw [e] at h' · exact let ⟨d, o, ad⟩ := IH _ _ h' - let ⟨c, cl, ac⟩ := this a ⟨d, o.resolve_right (not_mem_nil _), ad⟩ + let ⟨c, cl, ac⟩ := this a ⟨d, o.resolve_right (wseq.not_mem_nil _), ad⟩ ⟨c, Or.inl cl, ac⟩ · cases' a with o S' cases' o with c <;> simp [parallel.aux1] at h' <;> rcases IH _ _ h' with ⟨d, dl | dS', ad⟩ @@ -265,21 +263,21 @@ theorem exists_of_mem_parallel {S : Wseq (Computation α)} {a} (h : a ∈ parall let ⟨c, cl, ac⟩ := this a ⟨d, dl, ad⟩ ⟨c, Or.inl cl, ac⟩ · refine' ⟨d, Or.inr _, ad⟩ - rw [SeqCat.destruct_eq_cons e] - exact SeqCat.mem_cons_of_mem _ dS' + rw [seq.destruct_eq_cons e] + exact seq.mem_cons_of_mem _ dS' · simp at dl cases' dl with dc dl · rw [dc] at ad refine' ⟨c, Or.inr _, ad⟩ - rw [SeqCat.destruct_eq_cons e] - apply SeqCat.mem_cons + rw [seq.destruct_eq_cons e] + apply seq.mem_cons · exact let ⟨c, cl, ac⟩ := this a ⟨d, dl, ad⟩ ⟨c, Or.inl cl, ac⟩ · refine' ⟨d, Or.inr _, ad⟩ - rw [SeqCat.destruct_eq_cons e] - exact SeqCat.mem_cons_of_mem _ dS' + rw [seq.destruct_eq_cons e] + exact seq.mem_cons_of_mem _ dS' #align computation.exists_of_mem_parallel Computation.exists_of_mem_parallel theorem map_parallel (f : α → β) (S) : map f (parallel S) = parallel (S.map (map f)) := @@ -314,8 +312,8 @@ theorem parallel_empty (S : Wseq (Computation α)) (h : S.headI ~> none) : paral eq_empty_of_not_terminates fun ⟨⟨a, m⟩⟩ => by let ⟨c, cs, ac⟩ := exists_of_mem_parallel m - let ⟨n, nm⟩ := exists_nth_of_mem cs - let ⟨c', h'⟩ := head_some_of_nth_some nm + let ⟨n, nm⟩ := Wseq.exists_nth_of_mem cs + let ⟨c', h'⟩ := Wseq.head_some_of_nth_some nm injection h h' #align computation.parallel_empty Computation.parallel_empty @@ -323,11 +321,11 @@ theorem parallel_empty (S : Wseq (Computation α)) (h : S.headI ~> none) : paral def parallelRec {S : Wseq (Computation α)} (C : α → Sort v) (H : ∀ s ∈ S, ∀ a ∈ s, C a) {a} (h : a ∈ parallel S) : C a := by - let T : Wseq (Computation (α × Computation α)) := S.map fun c => c.map fun a => (a, c) + let T : wseq (Computation (α × Computation α)) := S.map fun c => c.map fun a => (a, c) have : S = T.map (map fun c => c.1) := by - rw [← Wseq.map_comp] - refine' (Wseq.map_id _).symm.trans (congr_arg (fun f => Wseq.map f S) _) + rw [← wseq.map_comp] + refine' (wseq.map_id _).symm.trans (congr_arg (fun f => wseq.map f S) _) funext c dsimp [id, Function.comp] rw [← map_comp] @@ -345,7 +343,7 @@ def parallelRec {S : Wseq (Computation α)} (C : α → Sort v) (H : ∀ s ∈ S dsimp at cd cases cd rcases exists_of_mem_parallel dT with ⟨d', dT', ad'⟩ - rcases Wseq.exists_of_mem_map dT' with ⟨c', cs', e'⟩ + rcases wseq.exists_of_mem_map dT' with ⟨c', cs', e'⟩ rw [← e'] at ad' rcases exists_of_mem_map ad' with ⟨a', ac', e'⟩ injection e' with i1 i2 @@ -388,14 +386,14 @@ theorem parallel_congr_left {S T : Wseq (Computation α)} {a} (h1 : ∀ s ∈ S, have aa := parallel_promises h1 h <;> rw [← aa] <;> rw [← aa] at h <;> exact let ⟨s, sS, as⟩ := exists_of_mem_parallel h - let ⟨t, tT, st⟩ := Wseq.exists_of_liftRel_left H sS + let ⟨t, tT, st⟩ := wseq.exists_of_lift_rel_left H sS let aT := (st _).1 as mem_parallel h2 tT aT, fun h => by have aa := parallel_promises h2 h <;> rw [← aa] <;> rw [← aa] at h <;> exact let ⟨s, sS, as⟩ := exists_of_mem_parallel h - let ⟨t, tT, st⟩ := Wseq.exists_of_liftRel_right H sS + let ⟨t, tT, st⟩ := wseq.exists_of_lift_rel_right H sS let aT := (st _).2 as mem_parallel h1 tT aT⟩ #align computation.parallel_congr_left Computation.parallel_congr_left diff --git a/Mathbin/Data/Seq/Seq.lean b/Mathbin/Data/Seq/Seq.lean index a61279357d..0b25db5308 100644 --- a/Mathbin/Data/Seq/Seq.lean +++ b/Mathbin/Data/Seq/Seq.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro ! This file was ported from Lean 3 source module data.seq.seq -! leanprover-community/mathlib commit 995b47e555f1b6297c7cf16855f1023e355219fb +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -14,6 +14,8 @@ import Mathbin.Data.Nat.Basic import Mathbin.Data.Stream.Init import Mathbin.Data.Seq.Computation +namespace Stream' + universe u v w /- @@ -23,187 +25,187 @@ coinductive seq (α : Type u) : Type u -/ /-- A stream `s : option α` is a sequence if `s.nth n = none` implies `s.nth (n + 1) = none`. -/ -def Stream'.IsSeq {α : Type u} (s : Stream' (Option α)) : Prop := +def IsSeq {α : Type u} (s : Stream' (Option α)) : Prop := ∀ {n : ℕ}, s n = none → s (n + 1) = none #align stream.is_seq Stream'.IsSeq /-- `seq α` is the type of possibly infinite lists (referred here as sequences). It is encoded as an infinite stream of options such that if `f n = none`, then `f m = none` for all `m ≥ n`. -/ -def SeqCat (α : Type u) : Type u := +def Seq (α : Type u) : Type u := { f : Stream' (Option α) // f.IsSeq } -#align seq SeqCat +#align stream.seq Stream'.Seq /-- `seq1 α` is the type of nonempty sequences. -/ def Seq1 (α) := - α × SeqCat α -#align seq1 Seq1 + α × Seq α +#align stream.seq1 Stream'.Seq1 -namespace SeqCat +namespace Seq variable {α : Type u} {β : Type v} {γ : Type w} /-- The empty sequence -/ -def nil : SeqCat α := +def nil : Seq α := ⟨Stream'.const none, fun n h => rfl⟩ -#align seq.nil SeqCat.nil +#align stream.seq.nil Stream'.Seq.nil -instance : Inhabited (SeqCat α) := +instance : Inhabited (Seq α) := ⟨nil⟩ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /-- Prepend an element to a sequence -/ -def cons (a : α) (s : SeqCat α) : SeqCat α := +def cons (a : α) (s : Seq α) : Seq α := ⟨some a::s.1, by rintro (n | _) h · contradiction · exact s.2 h⟩ -#align seq.cons SeqCat.cons +#align stream.seq.cons Stream'.Seq.cons /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] -theorem val_cons (s : SeqCat α) (x : α) : (cons x s).val = some x::s.val := +theorem val_cons (s : Seq α) (x : α) : (cons x s).val = some x::s.val := rfl -#align seq.val_cons SeqCat.val_cons +#align stream.seq.val_cons Stream'.Seq.val_cons /-- Get the nth element of a sequence (if it exists) -/ -def nth : SeqCat α → ℕ → Option α := +def nth : Seq α → ℕ → Option α := Subtype.val -#align seq.nth SeqCat.nth +#align stream.seq.nth Stream'.Seq.nth @[simp] theorem nth_mk (f hf) : @nth α ⟨f, hf⟩ = f := rfl -#align seq.nth_mk SeqCat.nth_mk +#align stream.seq.nth_mk Stream'.Seq.nth_mk @[simp] theorem nth_nil (n : ℕ) : (@nil α).get? n = none := rfl -#align seq.nth_nil SeqCat.nth_nil +#align stream.seq.nth_nil Stream'.Seq.nth_nil @[simp] -theorem nth_cons_zero (a : α) (s : SeqCat α) : (cons a s).get? 0 = some a := +theorem nth_cons_zero (a : α) (s : Seq α) : (cons a s).get? 0 = some a := rfl -#align seq.nth_cons_zero SeqCat.nth_cons_zero +#align stream.seq.nth_cons_zero Stream'.Seq.nth_cons_zero @[simp] -theorem nth_cons_succ (a : α) (s : SeqCat α) (n : ℕ) : (cons a s).get? (n + 1) = s.get? n := +theorem nth_cons_succ (a : α) (s : Seq α) (n : ℕ) : (cons a s).get? (n + 1) = s.get? n := rfl -#align seq.nth_cons_succ SeqCat.nth_cons_succ +#align stream.seq.nth_cons_succ Stream'.Seq.nth_cons_succ @[ext] -protected theorem ext {s t : SeqCat α} (h : ∀ n : ℕ, s.get? n = t.get? n) : s = t := +protected theorem ext {s t : Seq α} (h : ∀ n : ℕ, s.get? n = t.get? n) : s = t := Subtype.eq <| funext h -#align seq.ext SeqCat.ext +#align stream.seq.ext Stream'.Seq.ext -theorem cons_injective2 : Function.Injective2 (cons : α → SeqCat α → SeqCat α) := fun x y s t h => +theorem cons_injective2 : Function.Injective2 (cons : α → Seq α → Seq α) := fun x y s t h => ⟨by rw [← Option.some_inj, ← nth_cons_zero, h, nth_cons_zero], - SeqCat.ext fun n => by simp_rw [← nth_cons_succ x s n, h, nth_cons_succ]⟩ -#align seq.cons_injective2 SeqCat.cons_injective2 + Seq.ext fun n => by simp_rw [← nth_cons_succ x s n, h, nth_cons_succ]⟩ +#align stream.seq.cons_injective2 Stream'.Seq.cons_injective2 -theorem cons_left_injective (s : SeqCat α) : Function.Injective fun x => cons x s := +theorem cons_left_injective (s : Seq α) : Function.Injective fun x => cons x s := cons_injective2.left _ -#align seq.cons_left_injective SeqCat.cons_left_injective +#align stream.seq.cons_left_injective Stream'.Seq.cons_left_injective theorem cons_right_injective (x : α) : Function.Injective (cons x) := cons_injective2.right _ -#align seq.cons_right_injective SeqCat.cons_right_injective +#align stream.seq.cons_right_injective Stream'.Seq.cons_right_injective /-- A sequence has terminated at position `n` if the value at position `n` equals `none`. -/ -def TerminatedAt (s : SeqCat α) (n : ℕ) : Prop := +def TerminatedAt (s : Seq α) (n : ℕ) : Prop := s.get? n = none -#align seq.terminated_at SeqCat.TerminatedAt +#align stream.seq.terminated_at Stream'.Seq.TerminatedAt /-- It is decidable whether a sequence terminates at a given position. -/ -instance terminatedAtDecidable (s : SeqCat α) (n : ℕ) : Decidable (s.TerminatedAt n) := +instance terminatedAtDecidable (s : Seq α) (n : ℕ) : Decidable (s.TerminatedAt n) := decidable_of_iff' (s.get? n).isNone <| by unfold terminated_at <;> cases s.nth n <;> simp -#align seq.terminated_at_decidable SeqCat.terminatedAtDecidable +#align stream.seq.terminated_at_decidable Stream'.Seq.terminatedAtDecidable /-- A sequence terminates if there is some position `n` at which it has terminated. -/ -def Terminates (s : SeqCat α) : Prop := +def Terminates (s : Seq α) : Prop := ∃ n : ℕ, s.TerminatedAt n -#align seq.terminates SeqCat.Terminates +#align stream.seq.terminates Stream'.Seq.Terminates -theorem not_terminates_iff {s : SeqCat α} : ¬s.Terminates ↔ ∀ n, (s.get? n).isSome := by +theorem not_terminates_iff {s : Seq α} : ¬s.Terminates ↔ ∀ n, (s.get? n).isSome := by simp [terminates, terminated_at, ← Ne.def, Option.ne_none_iff_isSome] -#align seq.not_terminates_iff SeqCat.not_terminates_iff +#align stream.seq.not_terminates_iff Stream'.Seq.not_terminates_iff /-- Functorial action of the functor `option (α × _)` -/ @[simp] def omap (f : β → γ) : Option (α × β) → Option (α × γ) | none => none | some (a, b) => some (a, f b) -#align seq.omap SeqCat.omap +#align stream.seq.omap Stream'.Seq.omap /-- Get the first element of a sequence -/ -def head (s : SeqCat α) : Option α := +def head (s : Seq α) : Option α := nth s 0 -#align seq.head SeqCat.head +#align stream.seq.head Stream'.Seq.head /-- Get the tail of a sequence (or `nil` if the sequence is `nil`) -/ -def tail (s : SeqCat α) : SeqCat α := +def tail (s : Seq α) : Seq α := ⟨s.1.tail, fun n => by cases' s with f al exact al⟩ -#align seq.tail SeqCat.tail +#align stream.seq.tail Stream'.Seq.tail -protected def Mem (a : α) (s : SeqCat α) := +/-- member definition for `seq`-/ +protected def Mem (a : α) (s : Seq α) := some a ∈ s.1 -#align seq.mem SeqCat.Mem +#align stream.seq.mem Stream'.Seq.Mem -instance : Membership α (SeqCat α) := - ⟨SeqCat.Mem⟩ +instance : Membership α (Seq α) := + ⟨Seq.Mem⟩ -theorem le_stable (s : SeqCat α) {m n} (h : m ≤ n) : s.get? m = none → s.get? n = none := +theorem le_stable (s : Seq α) {m n} (h : m ≤ n) : s.get? m = none → s.get? n = none := by cases' s with f al induction' h with n h IH exacts[id, fun h2 => al (IH h2)] -#align seq.le_stable SeqCat.le_stable +#align stream.seq.le_stable Stream'.Seq.le_stable /-- If a sequence terminated at position `n`, it also terminated at `m ≥ n `. -/ -theorem terminated_stable : - ∀ (s : SeqCat α) {m n : ℕ}, m ≤ n → s.TerminatedAt m → s.TerminatedAt n := +theorem terminated_stable : ∀ (s : Seq α) {m n : ℕ}, m ≤ n → s.TerminatedAt m → s.TerminatedAt n := le_stable -#align seq.terminated_stable SeqCat.terminated_stable +#align stream.seq.terminated_stable Stream'.Seq.terminated_stable /-- If `s.nth n = some aₙ` for some value `aₙ`, then there is also some value `aₘ` such that `s.nth = some aₘ` for `m ≤ n`. -/ -theorem ge_stable (s : SeqCat α) {aₙ : α} {n m : ℕ} (m_le_n : m ≤ n) +theorem ge_stable (s : Seq α) {aₙ : α} {n m : ℕ} (m_le_n : m ≤ n) (s_nth_eq_some : s.get? n = some aₙ) : ∃ aₘ : α, s.get? m = some aₘ := have : s.get? n ≠ none := by simp [s_nth_eq_some] have : s.get? m ≠ none := mt (s.le_stable m_le_n) this Option.ne_none_iff_exists'.mp this -#align seq.ge_stable SeqCat.ge_stable +#align stream.seq.ge_stable Stream'.Seq.ge_stable theorem not_mem_nil (a : α) : a ∉ @nil α := fun ⟨n, (h : some a = none)⟩ => by injection h -#align seq.not_mem_nil SeqCat.not_mem_nil +#align stream.seq.not_mem_nil Stream'.Seq.not_mem_nil -theorem mem_cons (a : α) : ∀ s : SeqCat α, a ∈ cons a s +theorem mem_cons (a : α) : ∀ s : Seq α, a ∈ cons a s | ⟨f, al⟩ => Stream'.mem_cons (some a) _ -#align seq.mem_cons SeqCat.mem_cons +#align stream.seq.mem_cons Stream'.Seq.mem_cons -theorem mem_cons_of_mem (y : α) {a : α} : ∀ {s : SeqCat α}, a ∈ s → a ∈ cons y s +theorem mem_cons_of_mem (y : α) {a : α} : ∀ {s : Seq α}, a ∈ s → a ∈ cons y s | ⟨f, al⟩ => Stream'.mem_cons_of_mem (some y) -#align seq.mem_cons_of_mem SeqCat.mem_cons_of_mem +#align stream.seq.mem_cons_of_mem Stream'.Seq.mem_cons_of_mem -theorem eq_or_mem_of_mem_cons {a b : α} : ∀ {s : SeqCat α}, a ∈ cons b s → a = b ∨ a ∈ s +theorem eq_or_mem_of_mem_cons {a b : α} : ∀ {s : Seq α}, a ∈ cons b s → a = b ∨ a ∈ s | ⟨f, al⟩, h => (Stream'.eq_or_mem_of_mem_cons h).imp_left fun h => by injection h -#align seq.eq_or_mem_of_mem_cons SeqCat.eq_or_mem_of_mem_cons +#align stream.seq.eq_or_mem_of_mem_cons Stream'.Seq.eq_or_mem_of_mem_cons @[simp] -theorem mem_cons_iff {a b : α} {s : SeqCat α} : a ∈ cons b s ↔ a = b ∨ a ∈ s := +theorem mem_cons_iff {a b : α} {s : Seq α} : a ∈ cons b s ↔ a = b ∨ a ∈ s := ⟨eq_or_mem_of_mem_cons, by rintro (rfl | m) <;> [apply mem_cons, exact mem_cons_of_mem _ m]⟩ -#align seq.mem_cons_iff SeqCat.mem_cons_iff +#align stream.seq.mem_cons_iff Stream'.Seq.mem_cons_iff /-- Destructor for a sequence, resulting in either `none` (for `nil`) or `some (a, s)` (for `cons a s`). -/ -def destruct (s : SeqCat α) : Option (Seq1 α) := +def destruct (s : Seq α) : Option (Seq1 α) := (fun a' => (a', s.tail)) <$> nth s 0 -#align seq.destruct SeqCat.destruct +#align stream.seq.destruct Stream'.Seq.destruct -theorem destruct_eq_nil {s : SeqCat α} : destruct s = none → s = nil := +theorem destruct_eq_nil {s : Seq α} : destruct s = none → s = nil := by dsimp [destruct] induction' f0 : nth s 0 with <;> intro h @@ -212,9 +214,9 @@ theorem destruct_eq_nil {s : SeqCat α} : destruct s = none → s = nil := induction' n with n IH exacts[f0, s.2 IH] · contradiction -#align seq.destruct_eq_nil SeqCat.destruct_eq_nil +#align stream.seq.destruct_eq_nil Stream'.Seq.destruct_eq_nil -theorem destruct_eq_cons {s : SeqCat α} {a s'} : destruct s = some (a, s') → s = cons a s' := +theorem destruct_eq_cons {s : Seq α} {a s'} : destruct s = some (a, s') → s = cons a s' := by dsimp [destruct] induction' f0 : nth s 0 with a' <;> intro h @@ -227,12 +229,12 @@ theorem destruct_eq_cons {s : SeqCat α} {a s'} : destruct s = some (a, s') → rw [h1] at f0 rw [← f0] exact (Stream'.eta f).symm -#align seq.destruct_eq_cons SeqCat.destruct_eq_cons +#align stream.seq.destruct_eq_cons Stream'.Seq.destruct_eq_cons @[simp] -theorem destruct_nil : destruct (nil : SeqCat α) = none := +theorem destruct_nil : destruct (nil : Seq α) = none := rfl -#align seq.destruct_nil SeqCat.destruct_nil +#align stream.seq.destruct_nil Stream'.Seq.destruct_nil @[simp] theorem destruct_cons (a : α) : ∀ s, destruct (cons a s) = some (a, s) @@ -240,39 +242,39 @@ theorem destruct_cons (a : α) : ∀ s, destruct (cons a s) = some (a, s) unfold cons destruct Functor.map apply congr_arg fun s => some (a, s) apply Subtype.eq; dsimp [tail]; rw [Stream'.tail_cons] -#align seq.destruct_cons SeqCat.destruct_cons +#align stream.seq.destruct_cons Stream'.Seq.destruct_cons -theorem head_eq_destruct (s : SeqCat α) : head s = Prod.fst <$> destruct s := by +theorem head_eq_destruct (s : Seq α) : head s = Prod.fst <$> destruct s := by unfold destruct head <;> cases nth s 0 <;> rfl -#align seq.head_eq_destruct SeqCat.head_eq_destruct +#align stream.seq.head_eq_destruct Stream'.Seq.head_eq_destruct @[simp] -theorem head_nil : head (nil : SeqCat α) = none := +theorem head_nil : head (nil : Seq α) = none := rfl -#align seq.head_nil SeqCat.head_nil +#align stream.seq.head_nil Stream'.Seq.head_nil @[simp] theorem head_cons (a : α) (s) : head (cons a s) = some a := by rw [head_eq_destruct, destruct_cons] <;> rfl -#align seq.head_cons SeqCat.head_cons +#align stream.seq.head_cons Stream'.Seq.head_cons @[simp] -theorem tail_nil : tail (nil : SeqCat α) = nil := +theorem tail_nil : tail (nil : Seq α) = nil := rfl -#align seq.tail_nil SeqCat.tail_nil +#align stream.seq.tail_nil Stream'.Seq.tail_nil @[simp] theorem tail_cons (a : α) (s) : tail (cons a s) = s := by cases' s with f al <;> apply Subtype.eq <;> dsimp [tail, cons] <;> rw [Stream'.tail_cons] -#align seq.tail_cons SeqCat.tail_cons +#align stream.seq.tail_cons Stream'.Seq.tail_cons @[simp] -theorem nth_tail (s : SeqCat α) (n) : nth (tail s) n = nth s (n + 1) := +theorem nth_tail (s : Seq α) (n) : nth (tail s) n = nth s (n + 1) := rfl -#align seq.nth_tail SeqCat.nth_tail +#align stream.seq.nth_tail Stream'.Seq.nth_tail /-- Recursion principle for sequences, compare with `list.rec_on`. -/ -def recOn {C : SeqCat α → Sort v} (s : SeqCat α) (h1 : C nil) (h2 : ∀ x s, C (cons x s)) : C s := +def recOn {C : Seq α → Sort v} (s : Seq α) (h1 : C nil) (h2 : ∀ x s, C (cons x s)) : C s := by induction' H : destruct s with v v · rw [destruct_eq_nil H] @@ -280,9 +282,9 @@ def recOn {C : SeqCat α → Sort v} (s : SeqCat α) (h1 : C nil) (h2 : ∀ x s, · cases' v with a s' rw [destruct_eq_cons H] apply h2 -#align seq.rec_on SeqCat.recOn +#align stream.seq.rec_on Stream'.Seq.recOn -theorem mem_rec_on {C : SeqCat α → Prop} {a s} (M : a ∈ s) +theorem mem_rec_on {C : Seq α → Prop} {a s} (M : a ∈ s) (h1 : ∀ b s', a = b ∨ C s' → C (cons b s')) : C s := by cases' M with k e; unfold Stream'.nth at e @@ -299,19 +301,20 @@ theorem mem_rec_on {C : SeqCat α → Prop} {a s} (M : a ∈ s) · have h_eq : (cons b s').val (Nat.succ k) = s'.val k := by cases s' <;> rfl rw [h_eq] at e apply h1 _ _ (Or.inr (IH e)) -#align seq.mem_rec_on SeqCat.mem_rec_on +#align stream.seq.mem_rec_on Stream'.Seq.mem_rec_on +/-- Corecursor over pairs of `option` values-/ def Corec.f (f : β → Option (α × β)) : Option β → Option α × Option β | none => (none, none) | some b => match f b with | none => (none, none) | some (a, b') => (some a, some b') -#align seq.corec.F SeqCat.Corec.f +#align stream.seq.corec.F Stream'.Seq.Corec.f /-- Corecursor for `seq α` as a coinductive type. Iterates `f` to produce new elements of the sequence until `none` is obtained. -/ -def corec (f : β → Option (α × β)) (b : β) : SeqCat α := +def corec (f : β → Option (α × β)) (b : β) : Seq α := by refine' ⟨Stream'.corec' (corec.F f) (some b), fun n h => _⟩ rw [Stream'.corec'_eq] @@ -329,7 +332,7 @@ def corec (f : β → Option (α × β)) (b : β) : SeqCat α := contradiction · rw [Stream'.corec'_eq (corec.F f) (corec.F f o).2, Stream'.corec'_eq (corec.F f) o] exact IH (corec.F f o).2 -#align seq.corec SeqCat.corec +#align stream.seq.corec Stream'.Seq.corec @[simp] theorem corec_eq (f : β → Option (α × β)) (b : β) : destruct (corec f b) = omap (corec f) (f b) := @@ -344,32 +347,34 @@ theorem corec_eq (f : β → Option (α × β)) (b : β) : destruct (corec f b) dsimp [corec, tail] rw [Stream'.corec'_eq, Stream'.tail_cons] dsimp [corec.F]; rw [h]; rfl -#align seq.corec_eq SeqCat.corec_eq +#align stream.seq.corec_eq Stream'.Seq.corec_eq section Bisim -variable (R : SeqCat α → SeqCat α → Prop) +variable (R : Seq α → Seq α → Prop) -- mathport name: R local infixl:50 " ~ " => R +/-- Bisimilarity relation over `option` of `seq1 α`-/ def BisimO : Option (Seq1 α) → Option (Seq1 α) → Prop | none, none => True | some (a, s), some (a', s') => a = a' ∧ R s s' | _, _ => False -#align seq.bisim_o SeqCat.BisimO +#align stream.seq.bisim_o Stream'.Seq.BisimO attribute [simp] bisim_o +/-- a relation is bisimiar if it meets the `bisim_o` test-/ def IsBisimulation := ∀ ⦃s₁ s₂⦄, s₁ ~ s₂ → BisimO R (destruct s₁) (destruct s₂) -#align seq.is_bisimulation SeqCat.IsBisimulation +#align stream.seq.is_bisimulation Stream'.Seq.IsBisimulation -- If two streams are bisimilar, then they are equal theorem eq_of_bisim (bisim : IsBisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s₁ = s₂ := by apply Subtype.eq - apply Stream'.eq_of_bisim fun x y => ∃ s s' : SeqCat α, s.1 = x ∧ s'.1 = y ∧ R s s' + apply Stream'.eq_of_bisim fun x y => ∃ s s' : seq α, s.1 = x ∧ s'.1 = y ∧ R s s' dsimp [Stream'.IsBisimulation] intro t₁ t₂ e exact @@ -394,118 +399,118 @@ theorem eq_of_bisim (bisim : IsBisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s rw [h1] exact h2 exact ⟨s₁, s₂, rfl, rfl, r⟩ -#align seq.eq_of_bisim SeqCat.eq_of_bisim +#align stream.seq.eq_of_bisim Stream'.Seq.eq_of_bisim end Bisim theorem coinduction : - ∀ {s₁ s₂ : SeqCat α}, + ∀ {s₁ s₂ : Seq α}, head s₁ = head s₂ → - (∀ (β : Type u) (fr : SeqCat α → β), fr s₁ = fr s₂ → fr (tail s₁) = fr (tail s₂)) → s₁ = s₂ + (∀ (β : Type u) (fr : Seq α → β), fr s₁ = fr s₂ → fr (tail s₁) = fr (tail s₂)) → s₁ = s₂ | ⟨f₁, a₁⟩, ⟨f₂, a₂⟩, hh, ht => Subtype.eq (Stream'.coinduction hh fun β fr => ht β fun s => fr s.1) -#align seq.coinduction SeqCat.coinduction +#align stream.seq.coinduction Stream'.Seq.coinduction -theorem coinduction2 (s) (f g : SeqCat α → SeqCat β) +theorem coinduction2 (s) (f g : Seq α → Seq β) (H : ∀ s, - BisimO (fun s1 s2 : SeqCat β => ∃ s : SeqCat α, s1 = f s ∧ s2 = g s) (destruct (f s)) + BisimO (fun s1 s2 : Seq β => ∃ s : Seq α, s1 = f s ∧ s2 = g s) (destruct (f s)) (destruct (g s))) : f s = g s := by refine' eq_of_bisim (fun s1 s2 => ∃ s, s1 = f s ∧ s2 = g s) _ ⟨s, rfl, rfl⟩ intro s1 s2 h; rcases h with ⟨s, h1, h2⟩ rw [h1, h2]; apply H -#align seq.coinduction2 SeqCat.coinduction2 +#align stream.seq.coinduction2 Stream'.Seq.coinduction2 /-- Embed a list as a sequence -/ -def ofList (l : List α) : SeqCat α := +def ofList (l : List α) : Seq α := ⟨List.get? l, fun n h => by rw [List.get?_eq_none] at h⊢ exact h.trans (Nat.le_succ n)⟩ -#align seq.of_list SeqCat.ofList +#align stream.seq.of_list Stream'.Seq.ofList -instance coeList : Coe (List α) (SeqCat α) := +instance coeList : Coe (List α) (Seq α) := ⟨ofList⟩ -#align seq.coe_list SeqCat.coeList +#align stream.seq.coe_list Stream'.Seq.coeList @[simp] -theorem ofList_nil : ofList [] = (nil : SeqCat α) := +theorem ofList_nil : ofList [] = (nil : Seq α) := rfl -#align seq.of_list_nil SeqCat.ofList_nil +#align stream.seq.of_list_nil Stream'.Seq.ofList_nil @[simp] theorem ofList_nth (l : List α) (n : ℕ) : (ofList l).get? n = l.get? n := rfl -#align seq.of_list_nth SeqCat.ofList_nth +#align stream.seq.of_list_nth Stream'.Seq.ofList_nth /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] theorem ofList_cons (a : α) (l : List α) : ofList (a::l) = cons a (ofList l) := by ext1 (_ | n) <;> rfl -#align seq.of_list_cons SeqCat.ofList_cons +#align stream.seq.of_list_cons Stream'.Seq.ofList_cons /-- Embed an infinite stream as a sequence -/ -def ofStream (s : Stream' α) : SeqCat α := +def ofStream (s : Stream' α) : Seq α := ⟨s.map some, fun n h => by contradiction⟩ -#align seq.of_stream SeqCat.ofStream +#align stream.seq.of_stream Stream'.Seq.ofStream -instance coeStream : Coe (Stream' α) (SeqCat α) := +instance coeStream : Coe (Stream' α) (Seq α) := ⟨ofStream⟩ -#align seq.coe_stream SeqCat.coeStream +#align stream.seq.coe_stream Stream'.Seq.coeStream /-- Embed a `lazy_list α` as a sequence. Note that even though this is non-meta, it will produce infinite sequences if used with cyclic `lazy_list`s created by meta constructions. -/ -def ofLazyList : LazyList α → SeqCat α := +def ofLazyList : LazyList α → Seq α := corec fun l => match l with | LazyList.nil => none | LazyList.cons a l' => some (a, l' ()) -#align seq.of_lazy_list SeqCat.ofLazyList +#align stream.seq.of_lazy_list Stream'.Seq.ofLazyList -instance coeLazyList : Coe (LazyList α) (SeqCat α) := +instance coeLazyList : Coe (LazyList α) (Seq α) := ⟨ofLazyList⟩ -#align seq.coe_lazy_list SeqCat.coeLazyList +#align stream.seq.coe_lazy_list Stream'.Seq.coeLazyList /-- Translate a sequence into a `lazy_list`. Since `lazy_list` and `list` are isomorphic as non-meta types, this function is necessarily meta. -/ -unsafe def to_lazy_list : SeqCat α → LazyList α +unsafe def to_lazy_list : Seq α → LazyList α | s => match destruct s with | none => LazyList.nil | some (a, s') => LazyList.cons a (to_lazy_list s') -#align seq.to_lazy_list seq.to_lazy_list +#align stream.seq.to_lazy_list stream.seq.to_lazy_list /-- Translate a sequence to a list. This function will run forever if run on an infinite sequence. -/ -unsafe def force_to_list (s : SeqCat α) : List α := +unsafe def force_to_list (s : Seq α) : List α := (to_lazy_list s).toList -#align seq.force_to_list seq.force_to_list +#align stream.seq.force_to_list stream.seq.force_to_list /-- The sequence of natural numbers some 0, some 1, ... -/ -def nats : SeqCat ℕ := +def nats : Seq ℕ := Stream'.nats -#align seq.nats SeqCat.nats +#align stream.seq.nats Stream'.Seq.nats @[simp] theorem nats_nth (n : ℕ) : nats.get? n = some n := rfl -#align seq.nats_nth SeqCat.nats_nth +#align stream.seq.nats_nth Stream'.Seq.nats_nth /-- Append two sequences. If `s₁` is infinite, then `s₁ ++ s₂ = s₁`, otherwise it puts `s₂` at the location of the `nil` in `s₁`. -/ -def append (s₁ s₂ : SeqCat α) : SeqCat α := - @corec α (SeqCat α × SeqCat α) +def append (s₁ s₂ : Seq α) : Seq α := + @corec α (Seq α × Seq α) (fun ⟨s₁, s₂⟩ => match destruct s₁ with | none => omap (fun s₂ => (nil, s₂)) (destruct s₂) | some (a, s₁') => some (a, s₁', s₂)) (s₁, s₂) -#align seq.append SeqCat.append +#align stream.seq.append Stream'.Seq.append /-- Map a function over a sequence. -/ -def map (f : α → β) : SeqCat α → SeqCat β +def map (f : α → β) : Seq α → Seq β | ⟨s, al⟩ => ⟨s.map (Option.map f), fun n => by @@ -513,13 +518,13 @@ def map (f : α → β) : SeqCat α → SeqCat β induction' e : s n with <;> intro · rw [al e] assumption; · contradiction⟩ -#align seq.map SeqCat.map +#align stream.seq.map Stream'.Seq.map /-- Flatten a sequence of sequences. (It is required that the sequences be nonempty to ensure productivity; in the case of an infinite sequence of `nil`, the first element is never generated.) -/ -def join : SeqCat (Seq1 α) → SeqCat α := +def join : Seq (Seq1 α) → Seq α := corec fun S => match destruct S with | none => none @@ -529,28 +534,28 @@ def join : SeqCat (Seq1 α) → SeqCat α := match destruct s with | none => S' | some s' => cons s' S') -#align seq.join SeqCat.join +#align stream.seq.join Stream'.Seq.join /-- Remove the first `n` elements from the sequence. -/ -def drop (s : SeqCat α) : ℕ → SeqCat α +def drop (s : Seq α) : ℕ → Seq α | 0 => s | n + 1 => tail (drop n) -#align seq.drop SeqCat.drop +#align stream.seq.drop Stream'.Seq.drop attribute [simp] drop /-- Take the first `n` elements of the sequence (producing a list) -/ -def take : ℕ → SeqCat α → List α +def take : ℕ → Seq α → List α | 0, s => [] | n + 1, s => match destruct s with | none => [] | some (x, r) => List.cons x (take n r) -#align seq.take SeqCat.take +#align stream.seq.take Stream'.Seq.take /-- Split a sequence at `n`, producing a finite initial segment and an infinite tail. -/ -def splitAt : ℕ → SeqCat α → List α × SeqCat α +def splitAt : ℕ → Seq α → List α × Seq α | 0, s => ([], s) | n + 1, s => match destruct s with @@ -558,75 +563,75 @@ def splitAt : ℕ → SeqCat α → List α × SeqCat α | some (x, s') => let (l, r) := split_at n s' (List.cons x l, r) -#align seq.split_at SeqCat.splitAt +#align stream.seq.split_at Stream'.Seq.splitAt section ZipWith /-- Combine two sequences with a function -/ -def zipWith (f : α → β → γ) (s₁ : SeqCat α) (s₂ : SeqCat β) : SeqCat γ := +def zipWith (f : α → β → γ) (s₁ : Seq α) (s₂ : Seq β) : Seq γ := ⟨fun n => Option.map₂ f (s₁.get? n) (s₂.get? n), fun n hn => Option.map₂_eq_none_iff.2 <| (Option.map₂_eq_none_iff.1 hn).imp s₁.2 s₂.2⟩ -#align seq.zip_with SeqCat.zipWith +#align stream.seq.zip_with Stream'.Seq.zipWith -variable {s : SeqCat α} {s' : SeqCat β} {n : ℕ} +variable {s : Seq α} {s' : Seq β} {n : ℕ} @[simp] theorem nth_zipWith (f : α → β → γ) (s s' n) : (zipWith f s s').get? n = Option.map₂ f (s.get? n) (s'.get? n) := rfl -#align seq.nth_zip_with SeqCat.nth_zipWith +#align stream.seq.nth_zip_with Stream'.Seq.nth_zipWith end ZipWith /-- Pair two sequences into a sequence of pairs -/ -def zip : SeqCat α → SeqCat β → SeqCat (α × β) := +def zip : Seq α → Seq β → Seq (α × β) := zipWith Prod.mk -#align seq.zip SeqCat.zip +#align stream.seq.zip Stream'.Seq.zip -theorem nth_zip (s : SeqCat α) (t : SeqCat β) (n : ℕ) : +theorem nth_zip (s : Seq α) (t : Seq β) (n : ℕ) : nth (zip s t) n = Option.map₂ Prod.mk (nth s n) (nth t n) := nth_zipWith _ _ _ _ -#align seq.nth_zip SeqCat.nth_zip +#align stream.seq.nth_zip Stream'.Seq.nth_zip /-- Separate a sequence of pairs into two sequences -/ -def unzip (s : SeqCat (α × β)) : SeqCat α × SeqCat β := +def unzip (s : Seq (α × β)) : Seq α × Seq β := (map Prod.fst s, map Prod.snd s) -#align seq.unzip SeqCat.unzip +#align stream.seq.unzip Stream'.Seq.unzip /-- Enumerate a sequence by tagging each element with its index. -/ -def enum (s : SeqCat α) : SeqCat (ℕ × α) := - SeqCat.zip nats s -#align seq.enum SeqCat.enum +def enum (s : Seq α) : Seq (ℕ × α) := + Seq.zip nats s +#align stream.seq.enum Stream'.Seq.enum @[simp] -theorem nth_enum (s : SeqCat α) (n : ℕ) : nth (enum s) n = Option.map (Prod.mk n) (nth s n) := +theorem nth_enum (s : Seq α) (n : ℕ) : nth (enum s) n = Option.map (Prod.mk n) (nth s n) := nth_zip _ _ _ -#align seq.nth_enum SeqCat.nth_enum +#align stream.seq.nth_enum Stream'.Seq.nth_enum @[simp] -theorem enum_nil : enum (nil : SeqCat α) = nil := +theorem enum_nil : enum (nil : Seq α) = nil := rfl -#align seq.enum_nil SeqCat.enum_nil +#align stream.seq.enum_nil Stream'.Seq.enum_nil /-- Convert a sequence which is known to terminate into a list -/ -def toList (s : SeqCat α) (h : s.Terminates) : List α := +def toList (s : Seq α) (h : s.Terminates) : List α := take (Nat.find h) s -#align seq.to_list SeqCat.toList +#align stream.seq.to_list Stream'.Seq.toList /-- Convert a sequence which is known not to terminate into a stream -/ -def toStream (s : SeqCat α) (h : ¬s.Terminates) : Stream' α := fun n => +def toStream (s : Seq α) (h : ¬s.Terminates) : Stream' α := fun n => Option.get <| not_terminates_iff.1 h n -#align seq.to_stream SeqCat.toStream +#align stream.seq.to_stream Stream'.Seq.toStream /-- Convert a sequence into either a list or a stream depending on whether it is finite or infinite. (Without decidability of the infiniteness predicate, this is not constructively possible.) -/ -def toListOrStream (s : SeqCat α) [Decidable s.Terminates] : Sum (List α) (Stream' α) := +def toListOrStream (s : Seq α) [Decidable s.Terminates] : Sum (List α) (Stream' α) := if h : s.Terminates then Sum.inl (toList s h) else Sum.inr (toStream s h) -#align seq.to_list_or_stream SeqCat.toListOrStream +#align stream.seq.to_list_or_stream Stream'.Seq.toListOrStream @[simp] -theorem nil_append (s : SeqCat α) : append nil s = s := +theorem nil_append (s : Seq α) : append nil s = s := by apply coinduction2; intro s dsimp [append]; rw [corec_eq] @@ -636,7 +641,7 @@ theorem nil_append (s : SeqCat α) : append nil s = s := rw [destruct_cons] dsimp exact ⟨rfl, s, rfl, rfl⟩ -#align seq.nil_append SeqCat.nil_append +#align stream.seq.nil_append Stream'.Seq.nil_append @[simp] theorem cons_append (a : α) (s t) : append (cons a s) t = cons a (append s t) := @@ -644,10 +649,10 @@ theorem cons_append (a : α) (s t) : append (cons a s) t = cons a (append s t) : dsimp [append]; rw [corec_eq] dsimp [append]; rw [destruct_cons] dsimp [append]; rfl -#align seq.cons_append SeqCat.cons_append +#align stream.seq.cons_append Stream'.Seq.cons_append @[simp] -theorem append_nil (s : SeqCat α) : append s nil = s := +theorem append_nil (s : Seq α) : append s nil = s := by apply coinduction2 s; intro s apply rec_on s _ _ @@ -656,10 +661,10 @@ theorem append_nil (s : SeqCat α) : append s nil = s := rw [cons_append, destruct_cons, destruct_cons] dsimp exact ⟨rfl, s, rfl, rfl⟩ -#align seq.append_nil SeqCat.append_nil +#align stream.seq.append_nil Stream'.Seq.append_nil @[simp] -theorem append_assoc (s t u : SeqCat α) : append (append s t) u = append s (append t u) := +theorem append_assoc (s t u : Seq α) : append (append s t) u = append s (append t u) := by apply eq_of_bisim fun s1 s2 => ∃ s t u, s1 = append (append s t) u ∧ s2 = append s (append t u) · intro s1 s2 h @@ -676,37 +681,37 @@ theorem append_assoc (s t u : SeqCat α) : append (append s t) u = append s (app · intro x s exact ⟨s, t, u, rfl, rfl⟩ · exact ⟨s, t, u, rfl, rfl⟩ -#align seq.append_assoc SeqCat.append_assoc +#align stream.seq.append_assoc Stream'.Seq.append_assoc @[simp] theorem map_nil (f : α → β) : map f nil = nil := rfl -#align seq.map_nil SeqCat.map_nil +#align stream.seq.map_nil Stream'.Seq.map_nil @[simp] theorem map_cons (f : α → β) (a) : ∀ s, map f (cons a s) = cons (f a) (map f s) | ⟨s, al⟩ => by apply Subtype.eq <;> dsimp [cons, map] <;> rw [Stream'.map_cons] <;> rfl -#align seq.map_cons SeqCat.map_cons +#align stream.seq.map_cons Stream'.Seq.map_cons @[simp] -theorem map_id : ∀ s : SeqCat α, map id s = s +theorem map_id : ∀ s : Seq α, map id s = s | ⟨s, al⟩ => by apply Subtype.eq <;> dsimp [map] rw [Option.map_id, Stream'.map_id] <;> rfl -#align seq.map_id SeqCat.map_id +#align stream.seq.map_id Stream'.Seq.map_id @[simp] theorem map_tail (f : α → β) : ∀ s, map f (tail s) = tail (map f s) | ⟨s, al⟩ => by apply Subtype.eq <;> dsimp [tail, map] <;> rw [Stream'.map_tail] <;> rfl -#align seq.map_tail SeqCat.map_tail +#align stream.seq.map_tail Stream'.Seq.map_tail -theorem map_comp (f : α → β) (g : β → γ) : ∀ s : SeqCat α, map (g ∘ f) s = map g (map f s) +theorem map_comp (f : α → β) (g : β → γ) : ∀ s : Seq α, map (g ∘ f) s = map g (map f s) | ⟨s, al⟩ => by apply Subtype.eq <;> dsimp [map] rw [Stream'.map_map] apply congr_arg fun f : _ → Option γ => Stream'.map f s ext ⟨⟩ <;> rfl -#align seq.map_comp SeqCat.map_comp +#align stream.seq.map_comp Stream'.Seq.map_comp @[simp] theorem map_append (f : α → β) (s t) : map f (append s t) = append (map f s) (map f t) := @@ -724,34 +729,34 @@ theorem map_append (f : α → β) (s t) : map f (append s t) = append (map f s) refine' ⟨nil, t, _, _⟩ <;> simp · intro x s refine' ⟨s, t, rfl, rfl⟩ -#align seq.map_append SeqCat.map_append +#align stream.seq.map_append Stream'.Seq.map_append @[simp] theorem map_nth (f : α → β) : ∀ s n, nth (map f s) n = (nth s n).map f | ⟨s, al⟩, n => rfl -#align seq.map_nth SeqCat.map_nth +#align stream.seq.map_nth Stream'.Seq.map_nth -instance : Functor SeqCat where map := @map +instance : Functor Seq where map := @map -instance : LawfulFunctor SeqCat where +instance : LawfulFunctor Seq where id_map := @map_id comp_map := @map_comp @[simp] -theorem join_nil : join nil = (nil : SeqCat α) := +theorem join_nil : join nil = (nil : Seq α) := destruct_eq_nil rfl -#align seq.join_nil SeqCat.join_nil +#align stream.seq.join_nil Stream'.Seq.join_nil @[simp] theorem join_cons_nil (a : α) (S) : join (cons (a, nil) S) = cons a (join S) := destruct_eq_cons <| by simp [join] -#align seq.join_cons_nil SeqCat.join_cons_nil +#align stream.seq.join_cons_nil Stream'.Seq.join_cons_nil @[simp] theorem join_cons_cons (a b : α) (s S) : join (cons (a, cons b s) S) = cons a (join (cons (b, s) S)) := destruct_eq_cons <| by simp [join] -#align seq.join_cons_cons SeqCat.join_cons_cons +#align stream.seq.join_cons_cons Stream'.Seq.join_cons_cons @[simp] theorem join_cons (a : α) (s S) : join (cons (a, s) S) = cons a (append s (join S)) := @@ -774,10 +779,10 @@ theorem join_cons (a : α) (s S) : join (cons (a, s) S) = cons a (append s (join · intro x s simp refine' Or.inr ⟨x, s, S, rfl, rfl⟩ -#align seq.join_cons SeqCat.join_cons +#align stream.seq.join_cons Stream'.Seq.join_cons @[simp] -theorem join_append (S T : SeqCat (Seq1 α)) : join (append S T) = append (join S) (join T) := +theorem join_append (S T : Seq (Seq1 α)) : join (append S T) = append (join S) (join T) := by apply eq_of_bisim fun s1 s2 => @@ -799,65 +804,65 @@ theorem join_append (S T : SeqCat (Seq1 α)) : join (append S T) = append (join · intro x s exact ⟨s, S, T, rfl, rfl⟩ · refine' ⟨nil, S, T, _, _⟩ <;> simp -#align seq.join_append SeqCat.join_append +#align stream.seq.join_append Stream'.Seq.join_append /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] theorem ofStream_cons (a : α) (s) : ofStream (a::s) = cons a (ofStream s) := by apply Subtype.eq <;> simp [of_stream, cons] <;> rw [Stream'.map_cons] -#align seq.of_stream_cons SeqCat.ofStream_cons +#align stream.seq.of_stream_cons Stream'.Seq.ofStream_cons @[simp] theorem ofList_append (l l' : List α) : ofList (l ++ l') = append (ofList l) (ofList l') := by induction l <;> simp [*] -#align seq.of_list_append SeqCat.ofList_append +#align stream.seq.of_list_append Stream'.Seq.ofList_append @[simp] theorem ofStream_append (l : List α) (s : Stream' α) : ofStream (l ++ₛ s) = append (ofList l) (ofStream s) := by induction l <;> simp [*, Stream'.nil_append_stream, Stream'.cons_append_stream] -#align seq.of_stream_append SeqCat.ofStream_append +#align stream.seq.of_stream_append Stream'.Seq.ofStream_append /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /-- Convert a sequence into a list, embedded in a computation to allow for the possibility of infinite sequences (in which case the computation never returns anything). -/ -def toList' {α} (s : SeqCat α) : Computation (List α) := - @Computation.corec (List α) (List α × SeqCat α) +def toList' {α} (s : Seq α) : Computation (List α) := + @Computation.corec (List α) (List α × Seq α) (fun ⟨l, s⟩ => match destruct s with | none => Sum.inl l.reverse | some (a, s') => Sum.inr (a::l, s')) ([], s) -#align seq.to_list' SeqCat.toList' +#align stream.seq.to_list' Stream'.Seq.toList' -theorem dropn_add (s : SeqCat α) (m) : ∀ n, drop s (m + n) = drop (drop s m) n +theorem dropn_add (s : Seq α) (m) : ∀ n, drop s (m + n) = drop (drop s m) n | 0 => rfl | n + 1 => congr_arg tail (dropn_add n) -#align seq.dropn_add SeqCat.dropn_add +#align stream.seq.dropn_add Stream'.Seq.dropn_add -theorem dropn_tail (s : SeqCat α) (n) : drop (tail s) n = drop s (n + 1) := by +theorem dropn_tail (s : Seq α) (n) : drop (tail s) n = drop s (n + 1) := by rw [add_comm] <;> symm <;> apply dropn_add -#align seq.dropn_tail SeqCat.dropn_tail +#align stream.seq.dropn_tail Stream'.Seq.dropn_tail @[simp] -theorem head_dropn (s : SeqCat α) (n) : head (drop s n) = nth s n := +theorem head_dropn (s : Seq α) (n) : head (drop s n) = nth s n := by induction' n with n IH generalizing s; · rfl rw [Nat.succ_eq_add_one, ← nth_tail, ← dropn_tail]; apply IH -#align seq.head_dropn SeqCat.head_dropn +#align stream.seq.head_dropn Stream'.Seq.head_dropn -theorem mem_map (f : α → β) {a : α} : ∀ {s : SeqCat α}, a ∈ s → f a ∈ map f s +theorem mem_map (f : α → β) {a : α} : ∀ {s : Seq α}, a ∈ s → f a ∈ map f s | ⟨g, al⟩ => Stream'.mem_map (Option.map f) -#align seq.mem_map SeqCat.mem_map +#align stream.seq.mem_map Stream'.Seq.mem_map -theorem exists_of_mem_map {f} {b : β} : ∀ {s : SeqCat α}, b ∈ map f s → ∃ a, a ∈ s ∧ f a = b +theorem exists_of_mem_map {f} {b : β} : ∀ {s : Seq α}, b ∈ map f s → ∃ a, a ∈ s ∧ f a = b | ⟨g, al⟩, h => by let ⟨o, om, oe⟩ := Stream'.exists_of_mem_map h cases' o with a <;> injection oe with h' <;> exact ⟨a, om, h'⟩ -#align seq.exists_of_mem_map SeqCat.exists_of_mem_map +#align stream.seq.exists_of_mem_map Stream'.Seq.exists_of_mem_map -theorem of_mem_append {s₁ s₂ : SeqCat α} {a : α} (h : a ∈ append s₁ s₂) : a ∈ s₁ ∨ a ∈ s₂ := +theorem of_mem_append {s₁ s₂ : Seq α} {a : α} (h : a ∈ append s₁ s₂) : a ∈ s₁ ∨ a ∈ s₂ := by have := h; revert this generalize e : append s₁ s₂ = ss; intro h; revert s₁ @@ -873,71 +878,72 @@ theorem of_mem_append {s₁ s₂ : SeqCat α} {a : α} (h : a ∈ append s₁ s cases' o with e' IH · simp [i1, e'] · exact Or.imp_left (mem_cons_of_mem _) (IH m i2) -#align seq.of_mem_append SeqCat.of_mem_append +#align stream.seq.of_mem_append Stream'.Seq.of_mem_append -theorem mem_append_left {s₁ s₂ : SeqCat α} {a : α} (h : a ∈ s₁) : a ∈ append s₁ s₂ := by +theorem mem_append_left {s₁ s₂ : Seq α} {a : α} (h : a ∈ s₁) : a ∈ append s₁ s₂ := by apply mem_rec_on h <;> intros <;> simp [*] -#align seq.mem_append_left SeqCat.mem_append_left +#align stream.seq.mem_append_left Stream'.Seq.mem_append_left @[simp] -theorem enum_cons (s : SeqCat α) (x : α) : +theorem enum_cons (s : Seq α) (x : α) : enum (cons x s) = cons (0, x) (map (Prod.map Nat.succ id) (enum s)) := by ext ⟨n⟩ : 1 · simp · simp only [nth_enum, nth_cons_succ, map_nth, Option.map_map] congr -#align seq.enum_cons SeqCat.enum_cons +#align stream.seq.enum_cons Stream'.Seq.enum_cons -end SeqCat +end Seq namespace Seq1 variable {α : Type u} {β : Type v} {γ : Type w} -open SeqCat +open Stream'.Seq /-- Convert a `seq1` to a sequence. -/ -def toSeq : Seq1 α → SeqCat α - | (a, s) => cons a s -#align seq1.to_seq Seq1.toSeq +def toSeq : Seq1 α → Seq α + | (a, s) => Seq.cons a s +#align stream.seq1.to_seq Stream'.Seq1.toSeq -instance coeSeq : Coe (Seq1 α) (SeqCat α) := +instance coeSeq : Coe (Seq1 α) (Seq α) := ⟨toSeq⟩ -#align seq1.coe_seq Seq1.coeSeq +#align stream.seq1.coe_seq Stream'.Seq1.coeSeq /-- Map a function on a `seq1` -/ def map (f : α → β) : Seq1 α → Seq1 β - | (a, s) => (f a, SeqCat.map f s) -#align seq1.map Seq1.map + | (a, s) => (f a, Seq.map f s) +#align stream.seq1.map Stream'.Seq1.map theorem map_id : ∀ s : Seq1 α, map id s = s | ⟨a, s⟩ => by simp [map] -#align seq1.map_id Seq1.map_id +#align stream.seq1.map_id Stream'.Seq1.map_id /-- Flatten a nonempty sequence of nonempty sequences -/ def join : Seq1 (Seq1 α) → Seq1 α | ((a, s), S) => match destruct s with - | none => (a, SeqCat.join S) - | some s' => (a, SeqCat.join (cons s' S)) -#align seq1.join Seq1.join + | none => (a, Seq.join S) + | some s' => (a, Seq.join (Seq.cons s' S)) +#align stream.seq1.join Stream'.Seq1.join @[simp] -theorem join_nil (a : α) (S) : join ((a, nil), S) = (a, SeqCat.join S) := +theorem join_nil (a : α) (S) : join ((a, nil), S) = (a, Seq.join S) := rfl -#align seq1.join_nil Seq1.join_nil +#align stream.seq1.join_nil Stream'.Seq1.join_nil @[simp] -theorem join_cons (a b : α) (s S) : join ((a, cons b s), S) = (a, SeqCat.join (cons (b, s) S)) := by +theorem join_cons (a b : α) (s S) : + join ((a, Seq.cons b s), S) = (a, Seq.join (Seq.cons (b, s) S)) := by dsimp [join] <;> rw [destruct_cons] <;> rfl -#align seq1.join_cons Seq1.join_cons +#align stream.seq1.join_cons Stream'.Seq1.join_cons /-- The `return` operator for the `seq1` monad, which produces a singleton sequence. -/ def ret (a : α) : Seq1 α := (a, nil) -#align seq1.ret Seq1.ret +#align stream.seq1.ret Stream'.Seq1.ret instance [Inhabited α] : Inhabited (Seq1 α) := ⟨ret default⟩ @@ -948,19 +954,19 @@ instance [Inhabited α] : Inhabited (Seq1 α) := may already produce an infinite result.) -/ def bind (s : Seq1 α) (f : α → Seq1 β) : Seq1 β := join (map f s) -#align seq1.bind Seq1.bind +#align stream.seq1.bind Stream'.Seq1.bind @[simp] -theorem join_map_ret (s : SeqCat α) : SeqCat.join (SeqCat.map ret s) = s := by +theorem join_map_ret (s : Seq α) : Seq.join (Seq.map ret s) = s := by apply coinduction2 s <;> intro s <;> apply rec_on s <;> simp [ret] -#align seq1.join_map_ret Seq1.join_map_ret +#align stream.seq1.join_map_ret Stream'.Seq1.join_map_ret @[simp] theorem bind_ret (f : α → β) : ∀ s, bind s (ret ∘ f) = map f s | ⟨a, s⟩ => by dsimp [bind, map]; change fun x => ret (f x) with ret ∘ f rw [map_comp]; simp [Function.comp, ret] -#align seq1.bind_ret Seq1.bind_ret +#align stream.seq1.bind_ret Stream'.Seq1.bind_ret @[simp] theorem ret_bind (a : α) (f : α → Seq1 β) : bind (ret a) f = f a := @@ -968,17 +974,15 @@ theorem ret_bind (a : α) (f : α → Seq1 β) : bind (ret a) f = f a := simp [ret, bind, map] cases' f a with a s apply rec_on s <;> intros <;> simp -#align seq1.ret_bind Seq1.ret_bind +#align stream.seq1.ret_bind Stream'.Seq1.ret_bind @[simp] -theorem map_join' (f : α → β) (S) : - SeqCat.map f (SeqCat.join S) = SeqCat.join (SeqCat.map (map f) S) := +theorem map_join' (f : α → β) (S) : Seq.map f (Seq.join S) = Seq.join (Seq.map (map f) S) := by apply - eq_of_bisim fun s1 s2 => + seq.eq_of_bisim fun s1 s2 => ∃ s S, - s1 = append s (SeqCat.map f (SeqCat.join S)) ∧ - s2 = append s (SeqCat.join (SeqCat.map (map f) S)) + s1 = seq.append s (seq.map f (seq.join S)) ∧ s2 = append s (seq.join (seq.map (map f) S)) · intro s1 s2 h exact match s1, s2, h with @@ -991,22 +995,21 @@ theorem map_join' (f : α → β) (S) : · intro x s refine' ⟨s, S, rfl, rfl⟩ · refine' ⟨nil, S, _, _⟩ <;> simp -#align seq1.map_join' Seq1.map_join' +#align stream.seq1.map_join' Stream'.Seq1.map_join' @[simp] theorem map_join (f : α → β) : ∀ S, map f (join S) = join (map (map f) S) | ((a, s), S) => by apply rec_on s <;> intros <;> simp [map] -#align seq1.map_join Seq1.map_join +#align stream.seq1.map_join Stream'.Seq1.map_join @[simp] -theorem join_join (SS : SeqCat (Seq1 (Seq1 α))) : - SeqCat.join (SeqCat.join SS) = SeqCat.join (SeqCat.map join SS) := +theorem join_join (SS : Seq (Seq1 (Seq1 α))) : + Seq.join (Seq.join SS) = Seq.join (Seq.map join SS) := by apply - eq_of_bisim fun s1 s2 => + seq.eq_of_bisim fun s1 s2 => ∃ s SS, - s1 = SeqCat.append s (SeqCat.join (SeqCat.join SS)) ∧ - s2 = SeqCat.append s (SeqCat.join (SeqCat.map join SS)) + s1 = seq.append s (seq.join (seq.join SS)) ∧ s2 = seq.append s (seq.join (seq.map join SS)) · intro s1 s2 h exact match s1, s2, h with @@ -1018,11 +1021,11 @@ theorem join_join (SS : SeqCat (Seq1 (Seq1 α))) : apply rec_on s <;> simp · exact ⟨_, _, rfl, rfl⟩ · intro x s - refine' ⟨cons x (append s (SeqCat.join S)), SS, _, _⟩ <;> simp + refine' ⟨seq.cons x (append s (seq.join S)), SS, _, _⟩ <;> simp · intro x s exact ⟨s, SS, rfl, rfl⟩ · refine' ⟨nil, SS, _, _⟩ <;> simp -#align seq1.join_join Seq1.join_join +#align stream.seq1.join_join Stream'.Seq1.join_join @[simp] theorem bind_assoc (s : Seq1 α) (f : α → Seq1 β) (g : β → Seq1 γ) : @@ -1033,13 +1036,13 @@ theorem bind_assoc (s : Seq1 α) (f : α → Seq1 β) (g : β → Seq1 γ) : rw [← map_comp] change fun x => join (map g (f x)) with join ∘ map g ∘ f rw [map_comp _ join] - generalize SeqCat.map (map g ∘ f) s = SS + generalize seq.map (map g ∘ f) s = SS rcases map g (f a) with ⟨⟨a, s⟩, S⟩ apply rec_on s <;> intros <;> apply rec_on S <;> intros <;> simp · cases' x with x t apply rec_on t <;> intros <;> simp · cases' x_1 with y t <;> simp -#align seq1.bind_assoc Seq1.bind_assoc +#align stream.seq1.bind_assoc Stream'.Seq1.bind_assoc instance : Monad Seq1 where map := @map @@ -1054,3 +1057,5 @@ instance : LawfulMonad Seq1 where end Seq1 +end Stream' + diff --git a/Mathbin/Data/Seq/Wseq.lean b/Mathbin/Data/Seq/Wseq.lean index 5cac2cbc17..2f90141aa5 100644 --- a/Mathbin/Data/Seq/Wseq.lean +++ b/Mathbin/Data/Seq/Wseq.lean @@ -4,13 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro ! This file was ported from Lean 3 source module data.seq.wseq -! leanprover-community/mathlib commit 70fd9563a21e7b963887c9360bd29b2393e6225a +! leanprover-community/mathlib commit a7e36e48519ab281320c4d192da6a7b348ce40ad ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ import Mathbin.Data.List.Basic import Mathbin.Data.Seq.Seq +namespace Stream' + open Function universe u v w @@ -33,267 +35,268 @@ coinductive wseq (α : Type u) : Type u under most interesting computation patterns on infinite lists, but conversely it is difficult to extract elements from it. -/ def Wseq (α) := - SeqCat (Option α) -#align wseq Wseq + Seq (Option α) +#align stream.wseq Stream'.Wseq namespace Wseq variable {α : Type u} {β : Type v} {γ : Type w} /-- Turn a sequence into a weak sequence -/ -def ofSeq : SeqCat α → Wseq α := +def ofSeq : Seq α → Wseq α := (· <$> ·) some -#align wseq.of_seq Wseq.ofSeq +#align stream.wseq.of_seq Stream'.Wseq.ofSeq /-- Turn a list into a weak sequence -/ def ofList (l : List α) : Wseq α := ofSeq l -#align wseq.of_list Wseq.ofList +#align stream.wseq.of_list Stream'.Wseq.ofList /-- Turn a stream into a weak sequence -/ def ofStream (l : Stream' α) : Wseq α := ofSeq l -#align wseq.of_stream Wseq.ofStream +#align stream.wseq.of_stream Stream'.Wseq.ofStream -instance coeSeq : Coe (SeqCat α) (Wseq α) := +instance coeSeq : Coe (Seq α) (Wseq α) := ⟨ofSeq⟩ -#align wseq.coe_seq Wseq.coeSeq +#align stream.wseq.coe_seq Stream'.Wseq.coeSeq instance coeList : Coe (List α) (Wseq α) := ⟨ofList⟩ -#align wseq.coe_list Wseq.coeList +#align stream.wseq.coe_list Stream'.Wseq.coeList instance coeStream : Coe (Stream' α) (Wseq α) := ⟨ofStream⟩ -#align wseq.coe_stream Wseq.coeStream +#align stream.wseq.coe_stream Stream'.Wseq.coeStream /-- The empty weak sequence -/ def nil : Wseq α := - SeqCat.nil -#align wseq.nil Wseq.nil + Seq.nil +#align stream.wseq.nil Stream'.Wseq.nil instance : Inhabited (Wseq α) := ⟨nil⟩ /-- Prepend an element to a weak sequence -/ def cons (a : α) : Wseq α → Wseq α := - SeqCat.cons (some a) -#align wseq.cons Wseq.cons + Seq.cons (some a) +#align stream.wseq.cons Stream'.Wseq.cons /-- Compute for one tick, without producing any elements -/ def think : Wseq α → Wseq α := - SeqCat.cons none -#align wseq.think Wseq.think + Seq.cons none +#align stream.wseq.think Stream'.Wseq.think /-- Destruct a weak sequence, to (eventually possibly) produce either `none` for `nil` or `some (a, s)` if an element is produced. -/ def destruct : Wseq α → Computation (Option (α × Wseq α)) := Computation.corec fun s => - match SeqCat.destruct s with + match Seq.destruct s with | none => Sum.inl none | some (none, s') => Sum.inr s' | some (some a, s') => Sum.inl (some (a, s')) -#align wseq.destruct Wseq.destruct +#align stream.wseq.destruct Stream'.Wseq.destruct /-- Recursion principle for weak sequences, compare with `list.rec_on`. -/ def recOn {C : Wseq α → Sort v} (s : Wseq α) (h1 : C nil) (h2 : ∀ x s, C (cons x s)) (h3 : ∀ s, C (think s)) : C s := - SeqCat.recOn s h1 fun o => Option.recOn o h3 h2 -#align wseq.rec_on Wseq.recOn + Seq.recOn s h1 fun o => Option.recOn o h3 h2 +#align stream.wseq.rec_on Stream'.Wseq.recOn +/-- membership for weak sequences-/ protected def Mem (a : α) (s : Wseq α) := - SeqCat.Mem (some a) s -#align wseq.mem Wseq.Mem + Seq.Mem (some a) s +#align stream.wseq.mem Stream'.Wseq.Mem instance : Membership α (Wseq α) := ⟨Wseq.Mem⟩ theorem not_mem_nil (a : α) : a ∉ @nil α := - SeqCat.not_mem_nil a -#align wseq.not_mem_nil Wseq.not_mem_nil + Seq.not_mem_nil a +#align stream.wseq.not_mem_nil Stream'.Wseq.not_mem_nil /-- Get the head of a weak sequence. This involves a possibly infinite computation. -/ def head (s : Wseq α) : Computation (Option α) := Computation.map ((· <$> ·) Prod.fst) (destruct s) -#align wseq.head Wseq.head +#align stream.wseq.head Stream'.Wseq.head /-- Encode a computation yielding a weak sequence into additional `think` constructors in a weak sequence -/ def flatten : Computation (Wseq α) → Wseq α := - SeqCat.corec fun c => + Seq.corec fun c => match Computation.destruct c with - | Sum.inl s => SeqCat.omap return (SeqCat.destruct s) + | Sum.inl s => Seq.omap return (Seq.destruct s) | Sum.inr c' => some (none, c') -#align wseq.flatten Wseq.flatten +#align stream.wseq.flatten Stream'.Wseq.flatten /-- Get the tail of a weak sequence. This doesn't need a `computation` wrapper, unlike `head`, because `flatten` allows us to hide this in the construction of the weak sequence itself. -/ def tail (s : Wseq α) : Wseq α := flatten <| (fun o => Option.recOn o nil Prod.snd) <$> destruct s -#align wseq.tail Wseq.tail +#align stream.wseq.tail Stream'.Wseq.tail /-- drop the first `n` elements from `s`. -/ def drop (s : Wseq α) : ℕ → Wseq α | 0 => s | n + 1 => tail (drop n) -#align wseq.drop Wseq.drop +#align stream.wseq.drop Stream'.Wseq.drop attribute [simp] drop /-- Get the nth element of `s`. -/ def nth (s : Wseq α) (n : ℕ) : Computation (Option α) := head (drop s n) -#align wseq.nth Wseq.nth +#align stream.wseq.nth Stream'.Wseq.nth /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /-- Convert `s` to a list (if it is finite and completes in finite time). -/ def toList (s : Wseq α) : Computation (List α) := @Computation.corec (List α) (List α × Wseq α) (fun ⟨l, s⟩ => - match SeqCat.destruct s with + match Seq.destruct s with | none => Sum.inl l.reverse | some (none, s') => Sum.inr (l, s') | some (some a, s') => Sum.inr (a::l, s')) ([], s) -#align wseq.to_list Wseq.toList +#align stream.wseq.to_list Stream'.Wseq.toList /-- Get the length of `s` (if it is finite and completes in finite time). -/ def length (s : Wseq α) : Computation ℕ := @Computation.corec ℕ (ℕ × Wseq α) (fun ⟨n, s⟩ => - match SeqCat.destruct s with + match Seq.destruct s with | none => Sum.inl n | some (none, s') => Sum.inr (n, s') | some (some a, s') => Sum.inr (n + 1, s')) (0, s) -#align wseq.length Wseq.length +#align stream.wseq.length Stream'.Wseq.length /-- A weak sequence is finite if `to_list s` terminates. Equivalently, it is a finite number of `think` and `cons` applied to `nil`. -/ class IsFinite (s : Wseq α) : Prop where out : (toList s).Terminates -#align wseq.is_finite Wseq.IsFinite +#align stream.wseq.is_finite Stream'.Wseq.IsFinite instance toList_terminates (s : Wseq α) [h : IsFinite s] : (toList s).Terminates := h.out -#align wseq.to_list_terminates Wseq.toList_terminates +#align stream.wseq.to_list_terminates Stream'.Wseq.toList_terminates /-- Get the list corresponding to a finite weak sequence. -/ def get (s : Wseq α) [IsFinite s] : List α := (toList s).get -#align wseq.get Wseq.get +#align stream.wseq.get Stream'.Wseq.get /-- A weak sequence is *productive* if it never stalls forever - there are always a finite number of `think`s between `cons` constructors. The sequence itself is allowed to be infinite though. -/ class Productive (s : Wseq α) : Prop where nth_terminates : ∀ n, (nth s n).Terminates -#align wseq.productive Wseq.Productive +#align stream.wseq.productive Stream'.Wseq.Productive theorem productive_iff (s : Wseq α) : Productive s ↔ ∀ n, (nth s n).Terminates := ⟨fun h => h.1, fun h => ⟨h⟩⟩ -#align wseq.productive_iff Wseq.productive_iff +#align stream.wseq.productive_iff Stream'.Wseq.productive_iff instance nth_terminates (s : Wseq α) [h : Productive s] : ∀ n, (nth s n).Terminates := h.nth_terminates -#align wseq.nth_terminates Wseq.nth_terminates +#align stream.wseq.nth_terminates Stream'.Wseq.nth_terminates instance head_terminates (s : Wseq α) [Productive s] : (head s).Terminates := s.nth_terminates 0 -#align wseq.head_terminates Wseq.head_terminates +#align stream.wseq.head_terminates Stream'.Wseq.head_terminates /-- Replace the `n`th element of `s` with `a`. -/ def updateNth (s : Wseq α) (n : ℕ) (a : α) : Wseq α := - @SeqCat.corec (Option α) (ℕ × Wseq α) + @Seq.corec (Option α) (ℕ × Wseq α) (fun ⟨n, s⟩ => - match SeqCat.destruct s, n with + match Seq.destruct s, n with | none, n => none | some (none, s'), n => some (none, n, s') | some (some a', s'), 0 => some (some a', 0, s') | some (some a', s'), 1 => some (some a, 0, s') | some (some a', s'), n + 2 => some (some a', n + 1, s')) (n + 1, s) -#align wseq.update_nth Wseq.updateNth +#align stream.wseq.update_nth Stream'.Wseq.updateNth /-- Remove the `n`th element of `s`. -/ def removeNth (s : Wseq α) (n : ℕ) : Wseq α := - @SeqCat.corec (Option α) (ℕ × Wseq α) + @Seq.corec (Option α) (ℕ × Wseq α) (fun ⟨n, s⟩ => - match SeqCat.destruct s, n with + match Seq.destruct s, n with | none, n => none | some (none, s'), n => some (none, n, s') | some (some a', s'), 0 => some (some a', 0, s') | some (some a', s'), 1 => some (none, 0, s') | some (some a', s'), n + 2 => some (some a', n + 1, s')) (n + 1, s) -#align wseq.remove_nth Wseq.removeNth +#align stream.wseq.remove_nth Stream'.Wseq.removeNth /-- Map the elements of `s` over `f`, removing any values that yield `none`. -/ def filterMap (f : α → Option β) : Wseq α → Wseq β := - SeqCat.corec fun s => - match SeqCat.destruct s with + Seq.corec fun s => + match Seq.destruct s with | none => none | some (none, s') => some (none, s') | some (some a, s') => some (f a, s') -#align wseq.filter_map Wseq.filterMap +#align stream.wseq.filter_map Stream'.Wseq.filterMap /-- Select the elements of `s` that satisfy `p`. -/ def filter (p : α → Prop) [DecidablePred p] : Wseq α → Wseq α := filterMap fun a => if p a then some a else none -#align wseq.filter Wseq.filter +#align stream.wseq.filter Stream'.Wseq.filter -- example of infinite list manipulations /-- Get the first element of `s` satisfying `p`. -/ def find (p : α → Prop) [DecidablePred p] (s : Wseq α) : Computation (Option α) := head <| filter p s -#align wseq.find Wseq.find +#align stream.wseq.find Stream'.Wseq.find /-- Zip a function over two weak sequences -/ def zipWith (f : α → β → γ) (s1 : Wseq α) (s2 : Wseq β) : Wseq γ := - @SeqCat.corec (Option γ) (Wseq α × Wseq β) + @Seq.corec (Option γ) (Wseq α × Wseq β) (fun ⟨s1, s2⟩ => - match SeqCat.destruct s1, SeqCat.destruct s2 with + match Seq.destruct s1, Seq.destruct s2 with | some (none, s1'), some (none, s2') => some (none, s1', s2') | some (some a1, s1'), some (none, s2') => some (none, s1, s2') | some (none, s1'), some (some a2, s2') => some (none, s1', s2) | some (some a1, s1'), some (some a2, s2') => some (some (f a1 a2), s1', s2') | _, _ => none) (s1, s2) -#align wseq.zip_with Wseq.zipWith +#align stream.wseq.zip_with Stream'.Wseq.zipWith /-- Zip two weak sequences into a single sequence of pairs -/ def zip : Wseq α → Wseq β → Wseq (α × β) := zipWith Prod.mk -#align wseq.zip Wseq.zip +#align stream.wseq.zip Stream'.Wseq.zip /-- Get the list of indexes of elements of `s` satisfying `p` -/ def findIndexes (p : α → Prop) [DecidablePred p] (s : Wseq α) : Wseq ℕ := (zip s (Stream'.nats : Wseq ℕ)).filterMap fun ⟨a, n⟩ => if p a then some n else none -#align wseq.find_indexes Wseq.findIndexes +#align stream.wseq.find_indexes Stream'.Wseq.findIndexes /-- Get the index of the first element of `s` satisfying `p` -/ def findIndex (p : α → Prop) [DecidablePred p] (s : Wseq α) : Computation ℕ := (fun o => Option.getD o 0) <$> head (findIndexes p s) -#align wseq.find_index Wseq.findIndex +#align stream.wseq.find_index Stream'.Wseq.findIndex /-- Get the index of the first occurrence of `a` in `s` -/ def indexOf [DecidableEq α] (a : α) : Wseq α → Computation ℕ := findIndex (Eq a) -#align wseq.index_of Wseq.indexOf +#align stream.wseq.index_of Stream'.Wseq.indexOf /-- Get the indexes of occurrences of `a` in `s` -/ def indexesOf [DecidableEq α] (a : α) : Wseq α → Wseq ℕ := findIndexes (Eq a) -#align wseq.indexes_of Wseq.indexesOf +#align stream.wseq.indexes_of Stream'.Wseq.indexesOf /-- `union s1 s2` is a weak sequence which interleaves `s1` and `s2` in some order (nondeterministically). -/ def union (s1 s2 : Wseq α) : Wseq α := - @SeqCat.corec (Option α) (Wseq α × Wseq α) + @Seq.corec (Option α) (Wseq α × Wseq α) (fun ⟨s1, s2⟩ => - match SeqCat.destruct s1, SeqCat.destruct s2 with + match Seq.destruct s1, Seq.destruct s2 with | none, none => none | some (a1, s1'), none => some (a1, s1', nil) | none, some (a2, s2') => some (a2, nil, s2') @@ -302,31 +305,31 @@ def union (s1 s2 : Wseq α) : Wseq α := | some (none, s1'), some (some a2, s2') => some (some a2, s1', s2') | some (some a1, s1'), some (some a2, s2') => some (some a1, cons a2 s1', s2')) (s1, s2) -#align wseq.union Wseq.union +#align stream.wseq.union Stream'.Wseq.union /-- Returns `tt` if `s` is `nil` and `ff` if `s` has an element -/ def isEmpty (s : Wseq α) : Computation Bool := Computation.map Option.isNone <| head s -#align wseq.is_empty Wseq.isEmpty +#align stream.wseq.is_empty Stream'.Wseq.isEmpty /-- Calculate one step of computation -/ def compute (s : Wseq α) : Wseq α := - match SeqCat.destruct s with + match Seq.destruct s with | some (none, s') => s' | _ => s -#align wseq.compute Wseq.compute +#align stream.wseq.compute Stream'.Wseq.compute /-- Get the first `n` elements of a weak sequence -/ def take (s : Wseq α) (n : ℕ) : Wseq α := - @SeqCat.corec (Option α) (ℕ × Wseq α) + @Seq.corec (Option α) (ℕ × Wseq α) (fun ⟨n, s⟩ => - match n, SeqCat.destruct s with + match n, Seq.destruct s with | 0, _ => none | m + 1, none => none | m + 1, some (none, s') => some (none, m + 1, s') | m + 1, some (some a, s') => some (some a, m, s')) (n, s) -#align wseq.take Wseq.take +#align stream.wseq.take Stream'.Wseq.take /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /-- Split the sequence at position `n` into a finite initial segment @@ -334,106 +337,107 @@ def take (s : Wseq α) (n : ℕ) : Wseq α := def splitAt (s : Wseq α) (n : ℕ) : Computation (List α × Wseq α) := @Computation.corec (List α × Wseq α) (ℕ × List α × Wseq α) (fun ⟨n, l, s⟩ => - match n, SeqCat.destruct s with + match n, Seq.destruct s with | 0, _ => Sum.inl (l.reverse, s) | m + 1, none => Sum.inl (l.reverse, s) | m + 1, some (none, s') => Sum.inr (n, l, s') | m + 1, some (some a, s') => Sum.inr (m, a::l, s')) (n, [], s) -#align wseq.split_at Wseq.splitAt +#align stream.wseq.split_at Stream'.Wseq.splitAt /-- Returns `tt` if any element of `s` satisfies `p` -/ def any (s : Wseq α) (p : α → Bool) : Computation Bool := Computation.corec (fun s : Wseq α => - match SeqCat.destruct s with + match Seq.destruct s with | none => Sum.inl false | some (none, s') => Sum.inr s' | some (some a, s') => if p a then Sum.inl true else Sum.inr s') s -#align wseq.any Wseq.any +#align stream.wseq.any Stream'.Wseq.any /-- Returns `tt` if every element of `s` satisfies `p` -/ def all (s : Wseq α) (p : α → Bool) : Computation Bool := Computation.corec (fun s : Wseq α => - match SeqCat.destruct s with + match Seq.destruct s with | none => Sum.inl true | some (none, s') => Sum.inr s' | some (some a, s') => if p a then Sum.inr s' else Sum.inl false) s -#align wseq.all Wseq.all +#align stream.wseq.all Stream'.Wseq.all /-- Apply a function to the elements of the sequence to produce a sequence of partial results. (There is no `scanr` because this would require working from the end of the sequence, which may not exist.) -/ def scanl (f : α → β → α) (a : α) (s : Wseq β) : Wseq α := cons a <| - @SeqCat.corec (Option α) (α × Wseq β) + @Seq.corec (Option α) (α × Wseq β) (fun ⟨a, s⟩ => - match SeqCat.destruct s with + match Seq.destruct s with | none => none | some (none, s') => some (none, a, s') | some (some b, s') => let a' := f a b some (some a', a', s')) (a, s) -#align wseq.scanl Wseq.scanl +#align stream.wseq.scanl Stream'.Wseq.scanl /-- Get the weak sequence of initial segments of the input sequence -/ def inits (s : Wseq α) : Wseq (List α) := cons [] <| - @SeqCat.corec (Option (List α)) (Dlist α × Wseq α) + @Seq.corec (Option (List α)) (Dlist α × Wseq α) (fun ⟨l, s⟩ => - match SeqCat.destruct s with + match Seq.destruct s with | none => none | some (none, s') => some (none, l, s') | some (some a, s') => let l' := l.concat a some (some l'.toList, l', s')) (Dlist.empty, s) -#align wseq.inits Wseq.inits +#align stream.wseq.inits Stream'.Wseq.inits /-- Like take, but does not wait for a result. Calculates `n` steps of computation and returns the sequence computed so far -/ def collect (s : Wseq α) (n : ℕ) : List α := - (SeqCat.take n s).filterMap id -#align wseq.collect Wseq.collect + (Seq.take n s).filterMap id +#align stream.wseq.collect Stream'.Wseq.collect /-- Append two weak sequences. As with `seq.append`, this may not use the second sequence if the first one takes forever to compute -/ def append : Wseq α → Wseq α → Wseq α := - SeqCat.append -#align wseq.append Wseq.append + Seq.append +#align stream.wseq.append Stream'.Wseq.append /-- Map a function over a weak sequence -/ def map (f : α → β) : Wseq α → Wseq β := - SeqCat.map (Option.map f) -#align wseq.map Wseq.map + Seq.map (Option.map f) +#align stream.wseq.map Stream'.Wseq.map /-- Flatten a sequence of weak sequences. (Note that this allows empty sequences, unlike `seq.join`.) -/ def join (S : Wseq (Wseq α)) : Wseq α := - SeqCat.join + Seq.join ((fun o : Option (Wseq α) => match o with | none => Seq1.ret none | some s => (none, s)) <$> S) -#align wseq.join Wseq.join +#align stream.wseq.join Stream'.Wseq.join /-- Monadic bind operator for weak sequences -/ def bind (s : Wseq α) (f : α → Wseq β) : Wseq β := join (map f s) -#align wseq.bind Wseq.bind +#align stream.wseq.bind Stream'.Wseq.bind +/-- lift a relation to a relation over weak sequences -/ @[simp] def LiftRelO (R : α → β → Prop) (C : Wseq α → Wseq β → Prop) : Option (α × Wseq α) → Option (β × Wseq β) → Prop | none, none => True | some (a, s), some (b, t) => R a b ∧ C s t | _, _ => False -#align wseq.lift_rel_o Wseq.LiftRelO +#align stream.wseq.lift_rel_o Stream'.Wseq.LiftRelO theorem LiftRelO.imp {R S : α → β → Prop} {C D : Wseq α → Wseq β → Prop} (H1 : ∀ a b, R a b → S a b) (H2 : ∀ s t, C s t → D s t) : ∀ {o p}, LiftRelO R C o p → LiftRelO S D o p @@ -441,22 +445,23 @@ theorem LiftRelO.imp {R S : α → β → Prop} {C D : Wseq α → Wseq β → P | some (a, s), some (b, t), h => And.imp (H1 _ _) (H2 _ _) h | none, some _, h => False.elim h | some (_, _), none, h => False.elim h -#align wseq.lift_rel_o.imp Wseq.LiftRelO.imp +#align stream.wseq.lift_rel_o.imp Stream'.Wseq.LiftRelO.imp theorem LiftRelO.imp_right (R : α → β → Prop) {C D : Wseq α → Wseq β → Prop} (H : ∀ s t, C s t → D s t) {o p} : LiftRelO R C o p → LiftRelO R D o p := LiftRelO.imp (fun _ _ => id) H -#align wseq.lift_rel_o.imp_right Wseq.LiftRelO.imp_right +#align stream.wseq.lift_rel_o.imp_right Stream'.Wseq.LiftRelO.imp_right +/-- Definitino of bisimilarity for weak sequences-/ @[simp] def BisimO (R : Wseq α → Wseq α → Prop) : Option (α × Wseq α) → Option (α × Wseq α) → Prop := LiftRelO (· = ·) R -#align wseq.bisim_o Wseq.BisimO +#align stream.wseq.bisim_o Stream'.Wseq.BisimO theorem BisimO.imp {R S : Wseq α → Wseq α → Prop} (H : ∀ s t, R s t → S s t) {o p} : BisimO R o p → BisimO S o p := LiftRelO.imp_right _ H -#align wseq.bisim_o.imp Wseq.BisimO.imp +#align stream.wseq.bisim_o.imp Stream'.Wseq.BisimO.imp /-- Two weak sequences are `lift_rel R` related if they are either both empty, or they are both nonempty and the heads are `R` related and the tails are @@ -464,7 +469,7 @@ theorem BisimO.imp {R S : Wseq α → Wseq α → Prop} (H : ∀ s t, R s t → def LiftRel (R : α → β → Prop) (s : Wseq α) (t : Wseq β) : Prop := ∃ C : Wseq α → Wseq β → Prop, C s t ∧ ∀ {s t}, C s t → Computation.LiftRel (LiftRelO R C) (destruct s) (destruct t) -#align wseq.lift_rel Wseq.LiftRel +#align stream.wseq.lift_rel Stream'.Wseq.LiftRel /-- If two sequences are equivalent, then they have the same values and the same computational behavior (i.e. if one loops forever then so does @@ -472,14 +477,14 @@ def LiftRel (R : α → β → Prop) (s : Wseq α) (t : Wseq β) : Prop := arrive at the answer. -/ def Equiv : Wseq α → Wseq α → Prop := LiftRel (· = ·) -#align wseq.equiv Wseq.Equiv +#align stream.wseq.equiv Stream'.Wseq.Equiv theorem liftRel_destruct {R : α → β → Prop} {s : Wseq α} {t : Wseq β} : LiftRel R s t → Computation.LiftRel (LiftRelO R (LiftRel R)) (destruct s) (destruct t) | ⟨R, h1, h2⟩ => by refine' Computation.LiftRel.imp _ _ _ (h2 h1) <;> apply lift_rel_o.imp_right <;> exact fun s' t' h' => ⟨R, h', @h2⟩ -#align wseq.lift_rel_destruct Wseq.liftRel_destruct +#align stream.wseq.lift_rel_destruct Stream'.Wseq.liftRel_destruct theorem liftRel_destruct_iff {R : α → β → Prop} {s : Wseq α} {t : Wseq β} : LiftRel R s t ↔ Computation.LiftRel (LiftRelO R (LiftRel R)) (destruct s) (destruct t) := @@ -498,7 +503,7 @@ theorem liftRel_destruct_iff {R : α → β → Prop} {s : Wseq α} {t : Wseq β apply lift_rel_o.imp_right intro s t apply Or.inl⟩⟩ -#align wseq.lift_rel_destruct_iff Wseq.liftRel_destruct_iff +#align stream.wseq.lift_rel_destruct_iff Stream'.Wseq.liftRel_destruct_iff -- mathport name: equiv infixl:50 " ~ " => Equiv @@ -508,27 +513,27 @@ infixl:50 " ~ " => Equiv theorem destruct_congr {s t : Wseq α} : s ~ t → Computation.LiftRel (BisimO (· ~ ·)) (destruct s) (destruct t) := liftRel_destruct -#align wseq.destruct_congr Wseq.destruct_congr +#align stream.wseq.destruct_congr Stream'.Wseq.destruct_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:228:8: unsupported: ambiguous notation -/ theorem destruct_congr_iff {s t : Wseq α} : s ~ t ↔ Computation.LiftRel (BisimO (· ~ ·)) (destruct s) (destruct t) := liftRel_destruct_iff -#align wseq.destruct_congr_iff Wseq.destruct_congr_iff +#align stream.wseq.destruct_congr_iff Stream'.Wseq.destruct_congr_iff theorem LiftRel.refl (R : α → α → Prop) (H : Reflexive R) : Reflexive (LiftRel R) := fun s => by refine' ⟨(· = ·), rfl, fun s t (h : s = t) => _⟩ rw [← h]; apply Computation.LiftRel.refl intro a; cases' a with a; simp; cases a <;> simp; apply H -#align wseq.lift_rel.refl Wseq.LiftRel.refl +#align stream.wseq.lift_rel.refl Stream'.Wseq.LiftRel.refl theorem LiftRelO.swap (R : α → β → Prop) (C) : swap (LiftRelO R C) = LiftRelO (swap R) (swap C) := by funext x y <;> cases' x with x <;> [skip, cases x] <;> · cases' y with y <;> [skip, cases y] <;> rfl -#align wseq.lift_rel_o.swap Wseq.LiftRelO.swap +#align stream.wseq.lift_rel_o.swap Stream'.Wseq.LiftRelO.swap theorem LiftRel.swap_lem {R : α → β → Prop} {s1 s2} (h : LiftRel R s1 s2) : LiftRel (swap R) s2 s1 := @@ -536,17 +541,17 @@ theorem LiftRel.swap_lem {R : α → β → Prop} {s1 s2} (h : LiftRel R s1 s2) refine' ⟨swap (lift_rel R), h, fun s t (h : lift_rel R t s) => _⟩ rw [← lift_rel_o.swap, Computation.LiftRel.swap] apply lift_rel_destruct h -#align wseq.lift_rel.swap_lem Wseq.LiftRel.swap_lem +#align stream.wseq.lift_rel.swap_lem Stream'.Wseq.LiftRel.swap_lem theorem LiftRel.swap (R : α → β → Prop) : swap (LiftRel R) = LiftRel (swap R) := funext fun x => funext fun y => propext ⟨LiftRel.swap_lem, LiftRel.swap_lem⟩ -#align wseq.lift_rel.swap Wseq.LiftRel.swap +#align stream.wseq.lift_rel.swap Stream'.Wseq.LiftRel.swap theorem LiftRel.symm (R : α → α → Prop) (H : Symmetric R) : Symmetric (LiftRel R) := fun s1 s2 (h : swap (LiftRel R) s2 s1) => by rwa [lift_rel.swap, show swap R = R from funext fun a => funext fun b => propext <| by constructor <;> apply H] at h -#align wseq.lift_rel.symm Wseq.LiftRel.symm +#align stream.wseq.lift_rel.symm Stream'.Wseq.LiftRel.symm theorem LiftRel.trans (R : α → α → Prop) (H : Transitive R) : Transitive (LiftRel R) := fun s t u h1 h2 => @@ -579,24 +584,24 @@ theorem LiftRel.trans (R : α → α → Prop) (H : Transitive R) : Transitive ( cases' t1 with ab st cases' t2 with bc tu exact ⟨H ab bc, t, st, tu⟩ -#align wseq.lift_rel.trans Wseq.LiftRel.trans +#align stream.wseq.lift_rel.trans Stream'.Wseq.LiftRel.trans theorem LiftRel.equiv (R : α → α → Prop) : Equivalence R → Equivalence (LiftRel R) | ⟨refl, symm, trans⟩ => ⟨LiftRel.refl R refl, LiftRel.symm R symm, LiftRel.trans R trans⟩ -#align wseq.lift_rel.equiv Wseq.LiftRel.equiv +#align stream.wseq.lift_rel.equiv Stream'.Wseq.LiftRel.equiv /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[refl] theorem Equiv.refl : ∀ s : Wseq α, s ~ s := LiftRel.refl (· = ·) Eq.refl -#align wseq.equiv.refl Wseq.Equiv.refl +#align stream.wseq.equiv.refl Stream'.Wseq.Equiv.refl /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[symm] theorem Equiv.symm : ∀ {s t : Wseq α}, s ~ t → t ~ s := LiftRel.symm (· = ·) (@Eq.symm _) -#align wseq.equiv.symm Wseq.Equiv.symm +#align stream.wseq.equiv.symm Stream'.Wseq.Equiv.symm /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @@ -604,11 +609,11 @@ theorem Equiv.symm : ∀ {s t : Wseq α}, s ~ t → t ~ s := @[trans] theorem Equiv.trans : ∀ {s t u : Wseq α}, s ~ t → t ~ u → s ~ u := LiftRel.trans (· = ·) (@Eq.trans _) -#align wseq.equiv.trans Wseq.Equiv.trans +#align stream.wseq.equiv.trans Stream'.Wseq.Equiv.trans theorem Equiv.equivalence : Equivalence (@Equiv α) := ⟨@Equiv.refl _, @Equiv.symm _, @Equiv.trans _⟩ -#align wseq.equiv.equivalence Wseq.Equiv.equivalence +#align stream.wseq.equiv.equivalence Stream'.Wseq.Equiv.equivalence open Computation @@ -618,59 +623,59 @@ local notation "return" => Computation.pure @[simp] theorem destruct_nil : destruct (nil : Wseq α) = return none := Computation.destruct_eq_pure rfl -#align wseq.destruct_nil Wseq.destruct_nil +#align stream.wseq.destruct_nil Stream'.Wseq.destruct_nil @[simp] theorem destruct_cons (a : α) (s) : destruct (cons a s) = return (some (a, s)) := Computation.destruct_eq_pure <| by simp [destruct, cons, Computation.rmap] -#align wseq.destruct_cons Wseq.destruct_cons +#align stream.wseq.destruct_cons Stream'.Wseq.destruct_cons @[simp] theorem destruct_think (s : Wseq α) : destruct (think s) = (destruct s).think := Computation.destruct_eq_think <| by simp [destruct, think, Computation.rmap] -#align wseq.destruct_think Wseq.destruct_think +#align stream.wseq.destruct_think Stream'.Wseq.destruct_think @[simp] -theorem seq_destruct_nil : SeqCat.destruct (nil : Wseq α) = none := - SeqCat.destruct_nil -#align wseq.seq_destruct_nil Wseq.seq_destruct_nil +theorem seq_destruct_nil : Seq.destruct (nil : Wseq α) = none := + Seq.destruct_nil +#align stream.wseq.seq_destruct_nil Stream'.Wseq.seq_destruct_nil @[simp] -theorem seqCat_destruct_cons (a : α) (s) : SeqCat.destruct (cons a s) = some (some a, s) := - SeqCat.destruct_cons _ _ -#align wseq.seq_destruct_cons Wseq.seqCat_destruct_cons +theorem seq_destruct_cons (a : α) (s) : Seq.destruct (cons a s) = some (some a, s) := + Seq.destruct_cons _ _ +#align stream.wseq.seq_destruct_cons Stream'.Wseq.seq_destruct_cons @[simp] -theorem seqCat_destruct_think (s : Wseq α) : SeqCat.destruct (think s) = some (none, s) := - SeqCat.destruct_cons _ _ -#align wseq.seq_destruct_think Wseq.seqCat_destruct_think +theorem seq_destruct_think (s : Wseq α) : Seq.destruct (think s) = some (none, s) := + Seq.destruct_cons _ _ +#align stream.wseq.seq_destruct_think Stream'.Wseq.seq_destruct_think @[simp] theorem head_nil : head (nil : Wseq α) = return none := by simp [head] <;> rfl -#align wseq.head_nil Wseq.head_nil +#align stream.wseq.head_nil Stream'.Wseq.head_nil @[simp] theorem head_cons (a : α) (s) : head (cons a s) = return (some a) := by simp [head] <;> rfl -#align wseq.head_cons Wseq.head_cons +#align stream.wseq.head_cons Stream'.Wseq.head_cons @[simp] theorem head_think (s : Wseq α) : head (think s) = (head s).think := by simp [head] <;> rfl -#align wseq.head_think Wseq.head_think +#align stream.wseq.head_think Stream'.Wseq.head_think @[simp] theorem flatten_ret (s : Wseq α) : flatten (return s) = s := by - refine' SeqCat.eq_of_bisim (fun s1 s2 => flatten (return s2) = s1) _ rfl + refine' seq.eq_of_bisim (fun s1 s2 => flatten (return s2) = s1) _ rfl intro s' s h; rw [← h]; simp [flatten] - cases SeqCat.destruct s; · simp + cases seq.destruct s; · simp · cases' val with o s' simp -#align wseq.flatten_ret Wseq.flatten_ret +#align stream.wseq.flatten_ret Stream'.Wseq.flatten_ret @[simp] theorem flatten_think (c : Computation (Wseq α)) : flatten c.think = think (flatten c) := - SeqCat.destruct_eq_cons <| by simp [flatten, think] -#align wseq.flatten_think Wseq.flatten_think + Seq.destruct_eq_cons <| by simp [flatten, think] +#align stream.wseq.flatten_think Stream'.Wseq.flatten_think @[simp] theorem destruct_flatten (c : Computation (Wseq α)) : destruct (flatten c) = c >>= destruct := @@ -688,67 +693,67 @@ theorem destruct_flatten (c : Computation (Wseq α)) : destruct (flatten c) = c apply c.rec_on (fun a => _) fun c' => _ <;> repeat' simp · cases (destruct a).destruct <;> simp · exact Or.inr ⟨c', rfl, rfl⟩ -#align wseq.destruct_flatten Wseq.destruct_flatten +#align stream.wseq.destruct_flatten Stream'.Wseq.destruct_flatten theorem head_terminates_iff (s : Wseq α) : Terminates (head s) ↔ Terminates (destruct s) := terminates_map_iff _ (destruct s) -#align wseq.head_terminates_iff Wseq.head_terminates_iff +#align stream.wseq.head_terminates_iff Stream'.Wseq.head_terminates_iff @[simp] theorem tail_nil : tail (nil : Wseq α) = nil := by simp [tail] -#align wseq.tail_nil Wseq.tail_nil +#align stream.wseq.tail_nil Stream'.Wseq.tail_nil @[simp] theorem tail_cons (a : α) (s) : tail (cons a s) = s := by simp [tail] -#align wseq.tail_cons Wseq.tail_cons +#align stream.wseq.tail_cons Stream'.Wseq.tail_cons @[simp] theorem tail_think (s : Wseq α) : tail (think s) = (tail s).think := by simp [tail] -#align wseq.tail_think Wseq.tail_think +#align stream.wseq.tail_think Stream'.Wseq.tail_think @[simp] theorem dropn_nil (n) : drop (nil : Wseq α) n = nil := by induction n <;> simp [*, drop] -#align wseq.dropn_nil Wseq.dropn_nil +#align stream.wseq.dropn_nil Stream'.Wseq.dropn_nil @[simp] theorem dropn_cons (a : α) (s) (n) : drop (cons a s) (n + 1) = drop s n := by induction n <;> simp [*, drop] -#align wseq.dropn_cons Wseq.dropn_cons +#align stream.wseq.dropn_cons Stream'.Wseq.dropn_cons @[simp] theorem dropn_think (s : Wseq α) (n) : drop (think s) n = (drop s n).think := by induction n <;> simp [*, drop] -#align wseq.dropn_think Wseq.dropn_think +#align stream.wseq.dropn_think Stream'.Wseq.dropn_think theorem dropn_add (s : Wseq α) (m) : ∀ n, drop s (m + n) = drop (drop s m) n | 0 => rfl | n + 1 => congr_arg tail (dropn_add n) -#align wseq.dropn_add Wseq.dropn_add +#align stream.wseq.dropn_add Stream'.Wseq.dropn_add theorem dropn_tail (s : Wseq α) (n) : drop (tail s) n = drop s (n + 1) := by rw [add_comm] <;> symm <;> apply dropn_add -#align wseq.dropn_tail Wseq.dropn_tail +#align stream.wseq.dropn_tail Stream'.Wseq.dropn_tail theorem nth_add (s : Wseq α) (m n) : nth s (m + n) = nth (drop s m) n := congr_arg head (dropn_add _ _ _) -#align wseq.nth_add Wseq.nth_add +#align stream.wseq.nth_add Stream'.Wseq.nth_add theorem nth_tail (s : Wseq α) (n) : nth (tail s) n = nth s (n + 1) := congr_arg head (dropn_tail _ _) -#align wseq.nth_tail Wseq.nth_tail +#align stream.wseq.nth_tail Stream'.Wseq.nth_tail @[simp] theorem join_nil : join nil = (nil : Wseq α) := - SeqCat.join_nil -#align wseq.join_nil Wseq.join_nil + Seq.join_nil +#align stream.wseq.join_nil Stream'.Wseq.join_nil @[simp] theorem join_think (S : Wseq (Wseq α)) : join (think S) = think (join S) := by simp [think, join] unfold Functor.map - simp [join, Seq1.ret] -#align wseq.join_think Wseq.join_think + simp [join, seq1.ret] +#align stream.wseq.join_think Stream'.Wseq.join_think @[simp] theorem join_cons (s : Wseq α) (S) : join (cons s S) = think (append s (join S)) := @@ -756,62 +761,64 @@ theorem join_cons (s : Wseq α) (S) : join (cons s S) = think (append s (join S) simp [think, join] unfold Functor.map simp [join, cons, append] -#align wseq.join_cons Wseq.join_cons +#align stream.wseq.join_cons Stream'.Wseq.join_cons @[simp] theorem nil_append (s : Wseq α) : append nil s = s := - SeqCat.nil_append _ -#align wseq.nil_append Wseq.nil_append + Seq.nil_append _ +#align stream.wseq.nil_append Stream'.Wseq.nil_append @[simp] theorem cons_append (a : α) (s t) : append (cons a s) t = cons a (append s t) := - SeqCat.cons_append _ _ _ -#align wseq.cons_append Wseq.cons_append + Seq.cons_append _ _ _ +#align stream.wseq.cons_append Stream'.Wseq.cons_append @[simp] theorem think_append (s t : Wseq α) : append (think s) t = think (append s t) := - SeqCat.cons_append _ _ _ -#align wseq.think_append Wseq.think_append + Seq.cons_append _ _ _ +#align stream.wseq.think_append Stream'.Wseq.think_append @[simp] theorem append_nil (s : Wseq α) : append s nil = s := - SeqCat.append_nil _ -#align wseq.append_nil Wseq.append_nil + Seq.append_nil _ +#align stream.wseq.append_nil Stream'.Wseq.append_nil @[simp] theorem append_assoc (s t u : Wseq α) : append (append s t) u = append s (append t u) := - SeqCat.append_assoc _ _ _ -#align wseq.append_assoc Wseq.append_assoc + Seq.append_assoc _ _ _ +#align stream.wseq.append_assoc Stream'.Wseq.append_assoc +/-- auxilary defintion of tail over weak sequences-/ @[simp] def tail.aux : Option (α × Wseq α) → Computation (Option (α × Wseq α)) | none => return none | some (a, s) => destruct s -#align wseq.tail.aux Wseq.tail.aux +#align stream.wseq.tail.aux Stream'.Wseq.tail.aux theorem destruct_tail (s : Wseq α) : destruct (tail s) = destruct s >>= tail.aux := by simp [tail]; rw [← bind_pure_comp_eq_map, LawfulMonad.bind_assoc] apply congr_arg; ext1 (_ | ⟨a, s⟩) <;> apply (@pure_bind Computation _ _ _ _ _ _).trans _ <;> simp -#align wseq.destruct_tail Wseq.destruct_tail +#align stream.wseq.destruct_tail Stream'.Wseq.destruct_tail +/-- auxilary defintion of drop over weak sequences-/ @[simp] def drop.aux : ℕ → Option (α × Wseq α) → Computation (Option (α × Wseq α)) | 0 => return | n + 1 => fun a => tail.aux a >>= drop.aux n -#align wseq.drop.aux Wseq.drop.aux +#align stream.wseq.drop.aux Stream'.Wseq.drop.aux theorem drop.aux_none : ∀ n, @drop.aux α n none = return none | 0 => rfl | n + 1 => show Computation.bind (return none) (drop.aux n) = return none by rw [ret_bind, drop.aux_none] -#align wseq.drop.aux_none Wseq.drop.aux_none +#align stream.wseq.drop.aux_none Stream'.Wseq.drop.aux_none theorem destruct_dropn : ∀ (s : Wseq α) (n), destruct (drop s n) = destruct s >>= drop.aux n | s, 0 => (bind_pure' _).symm | s, n + 1 => by rw [← dropn_tail, destruct_dropn _ n, destruct_tail, LawfulMonad.bind_assoc] <;> rfl -#align wseq.destruct_dropn Wseq.destruct_dropn +#align stream.wseq.destruct_dropn Stream'.Wseq.destruct_dropn theorem head_terminates_of_head_tail_terminates (s : Wseq α) [T : Terminates (head (tail s))] : Terminates (head s) := @@ -822,29 +829,29 @@ theorem head_terminates_of_head_tail_terminates (s : Wseq α) [T : Terminates (h rcases exists_of_mem_bind h with ⟨s', h1, h2⟩ unfold Functor.map at h1 exact - let ⟨t, h3, h4⟩ := exists_of_mem_map h1 - terminates_of_mem h3 -#align wseq.head_terminates_of_head_tail_terminates Wseq.head_terminates_of_head_tail_terminates + let ⟨t, h3, h4⟩ := Computation.exists_of_mem_map h1 + Computation.terminates_of_mem h3 +#align stream.wseq.head_terminates_of_head_tail_terminates Stream'.Wseq.head_terminates_of_head_tail_terminates theorem destruct_some_of_destruct_tail_some {s : Wseq α} {a} (h : some a ∈ destruct (tail s)) : ∃ a', some a' ∈ destruct s := by unfold tail Functor.map at h; simp at h rcases exists_of_mem_bind h with ⟨t, tm, td⟩; clear h - rcases exists_of_mem_map tm with ⟨t', ht', ht2⟩; clear tm + rcases Computation.exists_of_mem_map tm with ⟨t', ht', ht2⟩; clear tm cases' t' with t' <;> rw [← ht2] at td <;> simp at td · have := mem_unique td (ret_mem _) contradiction · exact ⟨_, ht'⟩ -#align wseq.destruct_some_of_destruct_tail_some Wseq.destruct_some_of_destruct_tail_some +#align stream.wseq.destruct_some_of_destruct_tail_some Stream'.Wseq.destruct_some_of_destruct_tail_some theorem head_some_of_head_tail_some {s : Wseq α} {a} (h : some a ∈ head (tail s)) : ∃ a', some a' ∈ head s := by unfold head at h - rcases exists_of_mem_map h with ⟨o, md, e⟩; clear h + rcases Computation.exists_of_mem_map h with ⟨o, md, e⟩; clear h cases' o with o <;> injection e with h'; clear e h' cases' destruct_some_of_destruct_tail_some md with a am - exact ⟨_, mem_map ((· <$> ·) (@Prod.fst α (Wseq α))) am⟩ -#align wseq.head_some_of_head_tail_some Wseq.head_some_of_head_tail_some + exact ⟨_, Computation.mem_map ((· <$> ·) (@Prod.fst α (wseq α))) am⟩ +#align stream.wseq.head_some_of_head_tail_some Stream'.Wseq.head_some_of_head_tail_some theorem head_some_of_nth_some {s : Wseq α} {a n} (h : some a ∈ nth s n) : ∃ a', some a' ∈ head s := by @@ -852,19 +859,19 @@ theorem head_some_of_nth_some {s : Wseq α} {a n} (h : some a ∈ nth s n) : ∃ exacts[⟨_, h⟩, let ⟨a', h'⟩ := head_some_of_head_tail_some h IH h'] -#align wseq.head_some_of_nth_some Wseq.head_some_of_nth_some +#align stream.wseq.head_some_of_nth_some Stream'.Wseq.head_some_of_nth_some instance productive_tail (s : Wseq α) [Productive s] : Productive (tail s) := ⟨fun n => by rw [nth_tail] <;> infer_instance⟩ -#align wseq.productive_tail Wseq.productive_tail +#align stream.wseq.productive_tail Stream'.Wseq.productive_tail instance productive_dropn (s : Wseq α) [Productive s] (n) : Productive (drop s n) := ⟨fun m => by rw [← nth_add] <;> infer_instance⟩ -#align wseq.productive_dropn Wseq.productive_dropn +#align stream.wseq.productive_dropn Stream'.Wseq.productive_dropn /-- Given a productive weak sequence, we can collapse all the `think`s to produce a sequence. -/ -def toSeq (s : Wseq α) [Productive s] : SeqCat α := +def toSeq (s : Wseq α) [Productive s] : Seq α := ⟨fun n => (nth s n).get, fun n h => by cases e : Computation.get (nth s (n + 1)); · assumption @@ -872,28 +879,28 @@ def toSeq (s : Wseq α) [Productive s] : SeqCat α := simp [nth] at this h; cases' head_some_of_head_tail_some this with a' h' have := mem_unique h' (@mem_of_get_eq _ _ _ _ h) contradiction⟩ -#align wseq.to_seq Wseq.toSeq +#align stream.wseq.to_seq Stream'.Wseq.toSeq theorem nth_terminates_le {s : Wseq α} {m n} (h : m ≤ n) : Terminates (nth s n) → Terminates (nth s m) := by induction' h with m' h IH <;> [exact id, exact fun T => IH (@head_terminates_of_head_tail_terminates _ _ T)] -#align wseq.nth_terminates_le Wseq.nth_terminates_le +#align stream.wseq.nth_terminates_le Stream'.Wseq.nth_terminates_le theorem head_terminates_of_nth_terminates {s : Wseq α} {n} : Terminates (nth s n) → Terminates (head s) := nth_terminates_le (Nat.zero_le n) -#align wseq.head_terminates_of_nth_terminates Wseq.head_terminates_of_nth_terminates +#align stream.wseq.head_terminates_of_nth_terminates Stream'.Wseq.head_terminates_of_nth_terminates theorem destruct_terminates_of_nth_terminates {s : Wseq α} {n} (T : Terminates (nth s n)) : Terminates (destruct s) := (head_terminates_iff _).1 <| head_terminates_of_nth_terminates T -#align wseq.destruct_terminates_of_nth_terminates Wseq.destruct_terminates_of_nth_terminates +#align stream.wseq.destruct_terminates_of_nth_terminates Stream'.Wseq.destruct_terminates_of_nth_terminates theorem mem_rec_on {C : Wseq α → Prop} {a s} (M : a ∈ s) (h1 : ∀ b s', a = b ∨ C s' → C (cons b s')) (h2 : ∀ s, C s → C (think s)) : C s := by - apply SeqCat.mem_rec_on M + apply seq.mem_rec_on M intro o s' h; cases' o with b · apply h2 cases h @@ -903,7 +910,7 @@ theorem mem_rec_on {C : Wseq α → Prop} {a s} (M : a ∈ s) (h1 : ∀ b s', a apply Or.imp_left _ h intro h injection h -#align wseq.mem_rec_on Wseq.mem_rec_on +#align stream.wseq.mem_rec_on Stream'.Wseq.mem_rec_on /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] @@ -916,7 +923,7 @@ theorem mem_think (s : Wseq α) (a) : a ∈ think s ↔ a ∈ s := intro injections · apply Stream'.mem_cons_of_mem _ h -#align wseq.mem_think Wseq.mem_think +#align stream.wseq.mem_think Stream'.Wseq.mem_think theorem eq_or_mem_iff_mem {s : Wseq α} {a a' s'} : some (a', s') ∈ destruct s → (a ∈ s ↔ a = a' ∨ a ∈ s') := @@ -931,7 +938,7 @@ theorem eq_or_mem_iff_mem {s : Wseq α} {a a' s'} : cases' this with i1 i2 · rw [i1, i2] cases' s' with f al - unfold cons Membership.Mem Wseq.Mem SeqCat.Mem SeqCat.cons + unfold cons Membership.Mem wseq.mem seq.mem seq.cons simp have h_a_eq_a' : a = a' ↔ some (some a) = some (some a') := by simp rw [h_a_eq_a'] @@ -942,20 +949,20 @@ theorem eq_or_mem_iff_mem {s : Wseq α} {a a' s'} : · exact Stream'.mem_cons_of_mem _ m · simp exact IH this -#align wseq.eq_or_mem_iff_mem Wseq.eq_or_mem_iff_mem +#align stream.wseq.eq_or_mem_iff_mem Stream'.Wseq.eq_or_mem_iff_mem @[simp] theorem mem_cons_iff (s : Wseq α) (b) {a} : a ∈ cons b s ↔ a = b ∨ a ∈ s := eq_or_mem_iff_mem <| by simp [ret_mem] -#align wseq.mem_cons_iff Wseq.mem_cons_iff +#align stream.wseq.mem_cons_iff Stream'.Wseq.mem_cons_iff theorem mem_cons_of_mem {s : Wseq α} (b) {a} (h : a ∈ s) : a ∈ cons b s := (mem_cons_iff _ _).2 (Or.inr h) -#align wseq.mem_cons_of_mem Wseq.mem_cons_of_mem +#align stream.wseq.mem_cons_of_mem Stream'.Wseq.mem_cons_of_mem theorem mem_cons (s : Wseq α) (a) : a ∈ cons a s := (mem_cons_iff _ _).2 (Or.inl rfl) -#align wseq.mem_cons Wseq.mem_cons +#align stream.wseq.mem_cons Stream'.Wseq.mem_cons theorem mem_of_mem_tail {s : Wseq α} {a} : a ∈ tail s → a ∈ s := by @@ -970,24 +977,24 @@ theorem mem_of_mem_tail {s : Wseq α} {a} : a ∈ tail s → a ∈ s := rw [e] cases tail s rfl -#align wseq.mem_of_mem_tail Wseq.mem_of_mem_tail +#align stream.wseq.mem_of_mem_tail Stream'.Wseq.mem_of_mem_tail theorem mem_of_mem_dropn {s : Wseq α} {a} : ∀ {n}, a ∈ drop s n → a ∈ s | 0, h => h | n + 1, h => @mem_of_mem_dropn n (mem_of_mem_tail h) -#align wseq.mem_of_mem_dropn Wseq.mem_of_mem_dropn +#align stream.wseq.mem_of_mem_dropn Stream'.Wseq.mem_of_mem_dropn theorem nth_mem {s : Wseq α} {a n} : some a ∈ nth s n → a ∈ s := by revert s; induction' n with n IH <;> intro s h - · rcases exists_of_mem_map h with ⟨o, h1, h2⟩ + · rcases Computation.exists_of_mem_map h with ⟨o, h1, h2⟩ cases' o with o <;> injection h2 with h' cases' o with a' s' exact (eq_or_mem_iff_mem h1).2 (Or.inl h'.symm) · have := @IH (tail s) rw [nth_tail] at this exact mem_of_mem_tail (this h) -#align wseq.nth_mem Wseq.nth_mem +#align stream.wseq.nth_mem Stream'.Wseq.nth_mem theorem exists_nth_of_mem {s : Wseq α} {a} (h : a ∈ s) : ∃ n, some a ∈ nth s n := by @@ -1007,18 +1014,18 @@ theorem exists_nth_of_mem {s : Wseq α} {a} (h : a ∈ s) : ∃ n, some a ∈ nt exists n simp [nth] apply think_mem h -#align wseq.exists_nth_of_mem Wseq.exists_nth_of_mem +#align stream.wseq.exists_nth_of_mem Stream'.Wseq.exists_nth_of_mem theorem exists_dropn_of_mem {s : Wseq α} {a} (h : a ∈ s) : ∃ n s', some (a, s') ∈ destruct (drop s n) := let ⟨n, h⟩ := exists_nth_of_mem h ⟨n, by rcases(head_terminates_iff _).1 ⟨⟨_, h⟩⟩ with ⟨⟨o, om⟩⟩ - have := mem_unique (mem_map _ om) h + have := Computation.mem_unique (Computation.mem_map _ om) h cases' o with o <;> injection this with i cases' o with a' s'; dsimp at i rw [i] at om; exact ⟨_, om⟩⟩ -#align wseq.exists_dropn_of_mem Wseq.exists_dropn_of_mem +#align stream.wseq.exists_dropn_of_mem Stream'.Wseq.exists_dropn_of_mem theorem liftRel_dropn_destruct {R : α → β → Prop} {s t} (H : LiftRel R s t) : ∀ n, Computation.LiftRel (LiftRelO R (LiftRel R)) (destruct (drop s n)) (destruct (drop t n)) @@ -1031,74 +1038,74 @@ theorem liftRel_dropn_destruct {R : α → β → Prop} {s t} (H : LiftRel R s t match a, b, o with | none, none, _ => by simp | some (a, s), some (b, t), ⟨h1, h2⟩ => by simp [tail.aux] <;> apply lift_rel_destruct h2 -#align wseq.lift_rel_dropn_destruct Wseq.liftRel_dropn_destruct +#align stream.wseq.lift_rel_dropn_destruct Stream'.Wseq.liftRel_dropn_destruct theorem exists_of_liftRel_left {R : α → β → Prop} {s t} (H : LiftRel R s t) {a} (h : a ∈ s) : ∃ b, b ∈ t ∧ R a b := let ⟨n, h⟩ := exists_nth_of_mem h - let ⟨some (_, s'), sd, rfl⟩ := exists_of_mem_map h + let ⟨some (_, s'), sd, rfl⟩ := Computation.exists_of_mem_map h let ⟨some (b, t'), td, ⟨ab, _⟩⟩ := (liftRel_dropn_destruct H n).left sd - ⟨b, nth_mem (mem_map ((· <$> ·) Prod.fst.{v, v}) td), ab⟩ -#align wseq.exists_of_lift_rel_left Wseq.exists_of_liftRel_left + ⟨b, nth_mem (Computation.mem_map ((· <$> ·) Prod.fst.{v, v}) td), ab⟩ +#align stream.wseq.exists_of_lift_rel_left Stream'.Wseq.exists_of_liftRel_left theorem exists_of_liftRel_right {R : α → β → Prop} {s t} (H : LiftRel R s t) {b} (h : b ∈ t) : ∃ a, a ∈ s ∧ R a b := by rw [← lift_rel.swap] at H <;> exact exists_of_lift_rel_left H h -#align wseq.exists_of_lift_rel_right Wseq.exists_of_liftRel_right +#align stream.wseq.exists_of_lift_rel_right Stream'.Wseq.exists_of_liftRel_right theorem head_terminates_of_mem {s : Wseq α} {a} (h : a ∈ s) : Terminates (head s) := let ⟨n, h⟩ := exists_nth_of_mem h head_terminates_of_nth_terminates ⟨⟨_, h⟩⟩ -#align wseq.head_terminates_of_mem Wseq.head_terminates_of_mem +#align stream.wseq.head_terminates_of_mem Stream'.Wseq.head_terminates_of_mem theorem of_mem_append {s₁ s₂ : Wseq α} {a : α} : a ∈ append s₁ s₂ → a ∈ s₁ ∨ a ∈ s₂ := - SeqCat.of_mem_append -#align wseq.of_mem_append Wseq.of_mem_append + Seq.of_mem_append +#align stream.wseq.of_mem_append Stream'.Wseq.of_mem_append theorem mem_append_left {s₁ s₂ : Wseq α} {a : α} : a ∈ s₁ → a ∈ append s₁ s₂ := - SeqCat.mem_append_left -#align wseq.mem_append_left Wseq.mem_append_left + Seq.mem_append_left +#align stream.wseq.mem_append_left Stream'.Wseq.mem_append_left theorem exists_of_mem_map {f} {b : β} : ∀ {s : Wseq α}, b ∈ map f s → ∃ a, a ∈ s ∧ f a = b | ⟨g, al⟩, h => by - let ⟨o, om, oe⟩ := SeqCat.exists_of_mem_map h + let ⟨o, om, oe⟩ := Seq.exists_of_mem_map h cases' o with a <;> injection oe with h' <;> exact ⟨a, om, h'⟩ -#align wseq.exists_of_mem_map Wseq.exists_of_mem_map +#align stream.wseq.exists_of_mem_map Stream'.Wseq.exists_of_mem_map @[simp] theorem liftRel_nil (R : α → β → Prop) : LiftRel R nil nil := by rw [lift_rel_destruct_iff] <;> simp -#align wseq.lift_rel_nil Wseq.liftRel_nil +#align stream.wseq.lift_rel_nil Stream'.Wseq.liftRel_nil @[simp] theorem liftRel_cons (R : α → β → Prop) (a b s t) : LiftRel R (cons a s) (cons b t) ↔ R a b ∧ LiftRel R s t := by rw [lift_rel_destruct_iff] <;> simp -#align wseq.lift_rel_cons Wseq.liftRel_cons +#align stream.wseq.lift_rel_cons Stream'.Wseq.liftRel_cons @[simp] theorem liftRel_think_left (R : α → β → Prop) (s t) : LiftRel R (think s) t ↔ LiftRel R s t := by rw [lift_rel_destruct_iff, lift_rel_destruct_iff] <;> simp -#align wseq.lift_rel_think_left Wseq.liftRel_think_left +#align stream.wseq.lift_rel_think_left Stream'.Wseq.liftRel_think_left @[simp] theorem liftRel_think_right (R : α → β → Prop) (s t) : LiftRel R s (think t) ↔ LiftRel R s t := by rw [lift_rel_destruct_iff, lift_rel_destruct_iff] <;> simp -#align wseq.lift_rel_think_right Wseq.liftRel_think_right +#align stream.wseq.lift_rel_think_right Stream'.Wseq.liftRel_think_right /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem cons_congr {s t : Wseq α} (a : α) (h : s ~ t) : cons a s ~ cons a t := by unfold Equiv <;> simp <;> exact h -#align wseq.cons_congr Wseq.cons_congr +#align stream.wseq.cons_congr Stream'.Wseq.cons_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem think_equiv (s : Wseq α) : think s ~ s := by unfold Equiv <;> simp <;> apply Equiv.refl -#align wseq.think_equiv Wseq.think_equiv +#align stream.wseq.think_equiv Stream'.Wseq.think_equiv /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -theorem think_congr {s t : Wseq α} (a : α) (h : s ~ t) : think s ~ think t := by +theorem think_congr {s t : Wseq α} (h : s ~ t) : think s ~ think t := by unfold Equiv <;> simp <;> exact h -#align wseq.think_congr Wseq.think_congr +#align stream.wseq.think_congr Stream'.Wseq.think_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @@ -1113,7 +1120,7 @@ theorem head_congr : ∀ {s t : Wseq α}, s ~ t → head s ~ head t := cases' destruct_congr h with l r rcases l dsm with ⟨dt, dtm, dst⟩ cases' ds with a <;> cases' dt with b - · apply mem_map _ dtm + · apply Computation.mem_map _ dtm · cases b cases dst · cases a @@ -1121,8 +1128,8 @@ theorem head_congr : ∀ {s t : Wseq α}, s ~ t → head s ~ head t := · cases' a with a s' cases' b with b t' rw [dst.left] - exact @mem_map _ _ (@Functor.map _ _ (α × Wseq α) _ Prod.fst) _ (destruct t) dtm -#align wseq.head_congr Wseq.head_congr + exact @Computation.mem_map _ _ (@Functor.map _ _ (α × wseq α) _ Prod.fst) _ (destruct t) dtm +#align stream.wseq.head_congr Stream'.Wseq.head_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem flatten_equiv {c : Computation (Wseq α)} {s} (h : s ∈ c) : flatten c ~ s := @@ -1131,7 +1138,7 @@ theorem flatten_equiv {c : Computation (Wseq α)} {s} (h : s ∈ c) : flatten c · intro s' apply Equiv.trans simp [think_equiv] -#align wseq.flatten_equiv Wseq.flatten_equiv +#align stream.wseq.flatten_equiv Stream'.Wseq.flatten_equiv theorem liftRel_flatten {R : α → β → Prop} {c1 : Computation (Wseq α)} {c2 : Computation (Wseq β)} (h : c1.LiftRel (LiftRel R) c2) : LiftRel R (flatten c1) (flatten c2) := @@ -1143,13 +1150,13 @@ theorem liftRel_flatten {R : α → β → Prop} {c1 : Computation (Wseq α)} {c intro a b ab; apply Computation.LiftRel.imp _ _ _ (lift_rel_destruct ab) intro a b; apply lift_rel_o.imp_right intro s t h; refine' ⟨return s, return t, _, _, _⟩ <;> simp [h]⟩ -#align wseq.lift_rel_flatten Wseq.liftRel_flatten +#align stream.wseq.lift_rel_flatten Stream'.Wseq.liftRel_flatten /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem flatten_congr {c1 c2 : Computation (Wseq α)} : Computation.LiftRel Equiv c1 c2 → flatten c1 ~ flatten c2 := liftRel_flatten -#align wseq.flatten_congr Wseq.flatten_congr +#align stream.wseq.flatten_congr Stream'.Wseq.flatten_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @@ -1167,19 +1174,19 @@ theorem tail_congr {s t : Wseq α} (h : s ~ t) : tail s ~ tail t := · cases' a with a s' cases' b with b t' exact h.right -#align wseq.tail_congr Wseq.tail_congr +#align stream.wseq.tail_congr Stream'.Wseq.tail_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem dropn_congr {s t : Wseq α} (h : s ~ t) (n) : drop s n ~ drop t n := by induction n <;> simp [*, tail_congr] -#align wseq.dropn_congr Wseq.dropn_congr +#align stream.wseq.dropn_congr Stream'.Wseq.dropn_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem nth_congr {s t : Wseq α} (h : s ~ t) (n) : nth s n ~ nth t n := head_congr (dropn_congr h _) -#align wseq.nth_congr Wseq.nth_congr +#align stream.wseq.nth_congr Stream'.Wseq.nth_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @@ -1188,12 +1195,12 @@ theorem mem_congr {s t : Wseq α} (h : s ~ t) (a) : a ∈ s ↔ a ∈ t := fun s t h as => let ⟨n, hn⟩ := exists_nth_of_mem as nth_mem ((nth_congr h _ _).1 hn) -#align wseq.mem_congr Wseq.mem_congr +#align stream.wseq.mem_congr Stream'.Wseq.mem_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem productive_congr {s t : Wseq α} (h : s ~ t) : Productive s ↔ Productive t := by simp only [productive_iff] <;> exact forall_congr' fun n => terminates_congr <| nth_congr h _ -#align wseq.productive_congr Wseq.productive_congr +#align stream.wseq.productive_congr Stream'.Wseq.productive_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @@ -1208,188 +1215,187 @@ theorem Equiv.ext {s t : Wseq α} (h : ∀ n, nth s n ~ nth t n) : s ~ t := · intro a b ma mb cases' a with a <;> cases' b with b · trivial - · injection mem_unique (mem_map _ ma) ((h 0 _).2 (mem_map _ mb)) - · injection mem_unique (mem_map _ ma) ((h 0 _).2 (mem_map _ mb)) + · injection mem_unique (Computation.mem_map _ ma) ((h 0 _).2 (Computation.mem_map _ mb)) + · injection mem_unique (Computation.mem_map _ ma) ((h 0 _).2 (Computation.mem_map _ mb)) · cases' a with a s' cases' b with b t' - injection mem_unique (mem_map _ ma) ((h 0 _).2 (mem_map _ mb)) with ab + injection mem_unique (Computation.mem_map _ ma) ((h 0 _).2 (Computation.mem_map _ mb)) with + ab refine' ⟨ab, fun n => _⟩ refine' - (nth_congr (flatten_equiv (mem_map _ ma)) n).symm.trans + (nth_congr (flatten_equiv (Computation.mem_map _ ma)) n).symm.trans ((_ : nth (tail s) n ~ nth (tail t) n).trans - (nth_congr (flatten_equiv (mem_map _ mb)) n)) + (nth_congr (flatten_equiv (Computation.mem_map _ mb)) n)) rw [nth_tail, nth_tail] apply h⟩ -#align wseq.equiv.ext Wseq.Equiv.ext +#align stream.wseq.equiv.ext Stream'.Wseq.Equiv.ext /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem length_eq_map (s : Wseq α) : length s = Computation.map List.length (toList s) := by refine' - eq_of_bisim + Computation.eq_of_bisim (fun c1 c2 => - ∃ (l : List α)(s : Wseq α), - c1 = corec length._match_2 (l.length, s) ∧ - c2 = Computation.map List.length (corec to_list._match_2 (l, s))) + ∃ (l : List α)(s : wseq α), + c1 = Computation.corec length._match_2 (l.length, s) ∧ + c2 = Computation.map List.length (Computation.corec to_list._match_2 (l, s))) _ ⟨[], s, rfl, rfl⟩ intro s1 s2 h; rcases h with ⟨l, s, h⟩; rw [h.left, h.right] apply s.rec_on _ (fun a s => _) fun s => _ <;> repeat' simp [to_list, nil, cons, think, length] · refine' ⟨a::l, s, _, _⟩ <;> simp · refine' ⟨l, s, _, _⟩ <;> simp -#align wseq.length_eq_map Wseq.length_eq_map +#align stream.wseq.length_eq_map Stream'.Wseq.length_eq_map @[simp] theorem ofList_nil : ofList [] = (nil : Wseq α) := rfl -#align wseq.of_list_nil Wseq.ofList_nil +#align stream.wseq.of_list_nil Stream'.Wseq.ofList_nil /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] theorem ofList_cons (a : α) (l) : ofList (a::l) = cons a (ofList l) := - show - SeqCat.map some (SeqCat.ofList (a::l)) = - SeqCat.cons (some a) (SeqCat.map some (SeqCat.ofList l)) - by simp -#align wseq.of_list_cons Wseq.ofList_cons + show Seq.map some (Seq.ofList (a::l)) = Seq.cons (some a) (Seq.map some (Seq.ofList l)) by simp +#align stream.wseq.of_list_cons Stream'.Wseq.ofList_cons @[simp] -theorem to_list'_nil (l : List α) : corec ToList._match2 (l, nil) = return l.reverse := +theorem to_list'_nil (l : List α) : Computation.corec ToList._match2 (l, nil) = return l.reverse := destruct_eq_pure rfl -#align wseq.to_list'_nil Wseq.to_list'_nil +#align stream.wseq.to_list'_nil Stream'.Wseq.to_list'_nil /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] theorem to_list'_cons (l : List α) (s : Wseq α) (a : α) : - corec ToList._match2 (l, cons a s) = (corec ToList._match2 (a::l, s)).think := + Computation.corec ToList._match2 (l, cons a s) = + (Computation.corec ToList._match2 (a::l, s)).think := destruct_eq_think <| by simp [to_list, cons] -#align wseq.to_list'_cons Wseq.to_list'_cons +#align stream.wseq.to_list'_cons Stream'.Wseq.to_list'_cons @[simp] theorem to_list'_think (l : List α) (s : Wseq α) : - corec ToList._match2 (l, think s) = (corec ToList._match2 (l, s)).think := + Computation.corec ToList._match2 (l, think s) = + (Computation.corec ToList._match2 (l, s)).think := destruct_eq_think <| by simp [to_list, think] -#align wseq.to_list'_think Wseq.to_list'_think +#align stream.wseq.to_list'_think Stream'.Wseq.to_list'_think /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem to_list'_map (l : List α) (s : Wseq α) : - corec ToList._match2 (l, s) = (· ++ ·) l.reverse <$> toList s := + Computation.corec ToList._match2 (l, s) = (· ++ ·) l.reverse <$> toList s := by refine' - eq_of_bisim + Computation.eq_of_bisim (fun c1 c2 => - ∃ (l' : List α)(s : Wseq α), - c1 = corec to_list._match_2 (l' ++ l, s) ∧ - c2 = Computation.map ((· ++ ·) l.reverse) (corec to_list._match_2 (l', s))) + ∃ (l' : List α)(s : wseq α), + c1 = Computation.corec to_list._match_2 (l' ++ l, s) ∧ + c2 = Computation.map ((· ++ ·) l.reverse) (Computation.corec to_list._match_2 (l', s))) _ ⟨[], s, rfl, rfl⟩ intro s1 s2 h; rcases h with ⟨l', s, h⟩; rw [h.left, h.right] apply s.rec_on _ (fun a s => _) fun s => _ <;> repeat' simp [to_list, nil, cons, think, length] · refine' ⟨a::l', s, _, _⟩ <;> simp · refine' ⟨l', s, _, _⟩ <;> simp -#align wseq.to_list'_map Wseq.to_list'_map +#align stream.wseq.to_list'_map Stream'.Wseq.to_list'_map @[simp] theorem toList_cons (a : α) (s) : toList (cons a s) = (List.cons a <$> toList s).think := destruct_eq_think <| by unfold to_list <;> simp <;> rw [to_list'_map] <;> simp <;> rfl -#align wseq.to_list_cons Wseq.toList_cons +#align stream.wseq.to_list_cons Stream'.Wseq.toList_cons @[simp] theorem toList_nil : toList (nil : Wseq α) = return [] := destruct_eq_pure rfl -#align wseq.to_list_nil Wseq.toList_nil +#align stream.wseq.to_list_nil Stream'.Wseq.toList_nil theorem toList_ofList (l : List α) : l ∈ toList (ofList l) := by - induction' l with a l IH <;> simp [ret_mem] <;> exact think_mem (mem_map _ IH) -#align wseq.to_list_of_list Wseq.toList_ofList + induction' l with a l IH <;> simp [ret_mem] <;> exact think_mem (Computation.mem_map _ IH) +#align stream.wseq.to_list_of_list Stream'.Wseq.toList_ofList @[simp] -theorem destruct_ofSeq (s : SeqCat α) : +theorem destruct_ofSeq (s : Seq α) : destruct (ofSeq s) = return (s.headI.map fun a => (a, ofSeq s.tail)) := - destruct_eq_pure <| - by - simp [of_seq, head, destruct, SeqCat.destruct, SeqCat.head] - rw [show SeqCat.nth (some <$> s) 0 = some <$> SeqCat.nth s 0 by apply SeqCat.map_nth] - cases' SeqCat.nth s 0 with a; · rfl + destruct_eq_pure <| by + simp [of_seq, head, destruct, seq.destruct, seq.head] + rw [show seq.nth (some <$> s) 0 = some <$> seq.nth s 0 by apply seq.map_nth] + cases' seq.nth s 0 with a; · rfl unfold Functor.map simp [destruct] -#align wseq.destruct_of_seq Wseq.destruct_ofSeq +#align stream.wseq.destruct_of_seq Stream'.Wseq.destruct_ofSeq @[simp] -theorem head_ofSeq (s : SeqCat α) : head (ofSeq s) = return s.headI := by - simp [head] <;> cases SeqCat.head s <;> rfl -#align wseq.head_of_seq Wseq.head_ofSeq +theorem head_ofSeq (s : Seq α) : head (ofSeq s) = return s.headI := by + simp [head] <;> cases seq.head s <;> rfl +#align stream.wseq.head_of_seq Stream'.Wseq.head_ofSeq @[simp] -theorem tail_ofSeq (s : SeqCat α) : tail (ofSeq s) = ofSeq s.tail := +theorem tail_ofSeq (s : Seq α) : tail (ofSeq s) = ofSeq s.tail := by simp [tail]; apply s.rec_on _ fun x s => _ <;> simp [of_seq]; · rfl - rw [SeqCat.head_cons, SeqCat.tail_cons]; rfl -#align wseq.tail_of_seq Wseq.tail_ofSeq + rw [seq.head_cons, seq.tail_cons]; rfl +#align stream.wseq.tail_of_seq Stream'.Wseq.tail_ofSeq @[simp] -theorem dropn_ofSeq (s : SeqCat α) : ∀ n, drop (ofSeq s) n = ofSeq (s.drop n) +theorem dropn_ofSeq (s : Seq α) : ∀ n, drop (ofSeq s) n = ofSeq (s.drop n) | 0 => rfl | n + 1 => by dsimp [drop] <;> rw [dropn_of_seq, tail_of_seq] -#align wseq.dropn_of_seq Wseq.dropn_ofSeq +#align stream.wseq.dropn_of_seq Stream'.Wseq.dropn_ofSeq -theorem nth_ofSeq (s : SeqCat α) (n) : nth (ofSeq s) n = return (SeqCat.nth s n) := by - dsimp [nth] <;> rw [dropn_of_seq, head_of_seq, SeqCat.head_dropn] -#align wseq.nth_of_seq Wseq.nth_ofSeq +theorem nth_ofSeq (s : Seq α) (n) : nth (ofSeq s) n = return (Seq.nth s n) := by + dsimp [nth] <;> rw [dropn_of_seq, head_of_seq, seq.head_dropn] +#align stream.wseq.nth_of_seq Stream'.Wseq.nth_ofSeq -instance productive_ofSeq (s : SeqCat α) : Productive (ofSeq s) := +instance productive_ofSeq (s : Seq α) : Productive (ofSeq s) := ⟨fun n => by rw [nth_of_seq] <;> infer_instance⟩ -#align wseq.productive_of_seq Wseq.productive_ofSeq +#align stream.wseq.productive_of_seq Stream'.Wseq.productive_ofSeq -theorem toSeq_ofSeq (s : SeqCat α) : toSeq (ofSeq s) = s := +theorem toSeq_ofSeq (s : Seq α) : toSeq (ofSeq s) = s := by apply Subtype.eq; funext n dsimp [to_seq]; apply get_eq_of_mem rw [nth_of_seq]; apply ret_mem -#align wseq.to_seq_of_seq Wseq.toSeq_ofSeq +#align stream.wseq.to_seq_of_seq Stream'.Wseq.toSeq_ofSeq /-- The monadic `return a` is a singleton list containing `a`. -/ def ret (a : α) : Wseq α := ofList [a] -#align wseq.ret Wseq.ret +#align stream.wseq.ret Stream'.Wseq.ret @[simp] theorem map_nil (f : α → β) : map f nil = nil := rfl -#align wseq.map_nil Wseq.map_nil +#align stream.wseq.map_nil Stream'.Wseq.map_nil @[simp] theorem map_cons (f : α → β) (a s) : map f (cons a s) = cons (f a) (map f s) := - SeqCat.map_cons _ _ _ -#align wseq.map_cons Wseq.map_cons + Seq.map_cons _ _ _ +#align stream.wseq.map_cons Stream'.Wseq.map_cons @[simp] theorem map_think (f : α → β) (s) : map f (think s) = think (map f s) := - SeqCat.map_cons _ _ _ -#align wseq.map_think Wseq.map_think + Seq.map_cons _ _ _ +#align stream.wseq.map_think Stream'.Wseq.map_think @[simp] theorem map_id (s : Wseq α) : map id s = s := by simp [map] -#align wseq.map_id Wseq.map_id +#align stream.wseq.map_id Stream'.Wseq.map_id @[simp] theorem map_ret (f : α → β) (a) : map f (ret a) = ret (f a) := by simp [ret] -#align wseq.map_ret Wseq.map_ret +#align stream.wseq.map_ret Stream'.Wseq.map_ret @[simp] theorem map_append (f : α → β) (s t) : map f (append s t) = append (map f s) (map f t) := - SeqCat.map_append _ _ _ -#align wseq.map_append Wseq.map_append + Seq.map_append _ _ _ +#align stream.wseq.map_append Stream'.Wseq.map_append theorem map_comp (f : α → β) (g : β → γ) (s : Wseq α) : map (g ∘ f) s = map g (map f s) := by - dsimp [map]; rw [← SeqCat.map_comp] + dsimp [map]; rw [← seq.map_comp] apply congr_fun; apply congr_arg ext ⟨⟩ <;> rfl -#align wseq.map_comp Wseq.map_comp +#align stream.wseq.map_comp Stream'.Wseq.map_comp theorem mem_map (f : α → β) {a : α} {s : Wseq α} : a ∈ s → f a ∈ map f s := - SeqCat.mem_map (Option.map f) -#align wseq.mem_map Wseq.mem_map + Seq.mem_map (Option.map f) +#align stream.wseq.mem_map Stream'.Wseq.mem_map -- The converse is not true without additional assumptions theorem exists_of_mem_join {a : α} : ∀ {S : Wseq (Wseq α)}, a ∈ join S → ∃ s, s ∈ S ∧ a ∈ s := @@ -1400,8 +1406,8 @@ theorem exists_of_mem_join {a : α} : ∀ {S : Wseq (Wseq α)}, a ∈ join S → from fun S h => (this _ h nil S (by simp) (by simp [h])).resolve_left (not_mem_nil _) intro ss h; apply mem_rec_on h (fun b ss o => _) fun ss IH => _ <;> intro s S · refine' s.rec_on (S.rec_on _ (fun s S => _) fun S => _) (fun b' s => _) fun s => _ <;> - intro ej m <;> simp at ej <;> have := congr_arg SeqCat.destruct ej <;> - simp at this <;> try cases this <;> try contradiction + intro ej m <;> simp at ej <;> have := congr_arg seq.destruct ej <;> simp at this <;> + try cases this <;> try contradiction substs b' ss simp at m⊢ cases' o with e IH @@ -1410,8 +1416,8 @@ theorem exists_of_mem_join {a : α} : ∀ {S : Wseq (Wseq α)}, a ∈ join S → · simp [e] exact Or.imp_left Or.inr (IH _ _ rfl m) · refine' s.rec_on (S.rec_on _ (fun s S => _) fun S => _) (fun b' s => _) fun s => _ <;> - intro ej m <;> simp at ej <;> have := congr_arg SeqCat.destruct ej <;> - simp at this <;> try try have := this.1; contradiction <;> subst ss + intro ej m <;> simp at ej <;> have := congr_arg seq.destruct ej <;> simp at this <;> + try try have := this.1; contradiction <;> subst ss · apply Or.inr simp at m⊢ cases' IH s S rfl m with as ex @@ -1424,20 +1430,20 @@ theorem exists_of_mem_join {a : α} : ∀ {S : Wseq (Wseq α)}, a ∈ join S → exact ⟨s, by simp [sS], as⟩ · simp at m IH⊢ apply IH _ _ rfl m -#align wseq.exists_of_mem_join Wseq.exists_of_mem_join +#align stream.wseq.exists_of_mem_join Stream'.Wseq.exists_of_mem_join theorem exists_of_mem_bind {s : Wseq α} {f : α → Wseq β} {b} (h : b ∈ bind s f) : ∃ a ∈ s, b ∈ f a := let ⟨t, tm, bt⟩ := exists_of_mem_join h let ⟨a, as, e⟩ := exists_of_mem_map tm ⟨a, as, by rwa [e]⟩ -#align wseq.exists_of_mem_bind Wseq.exists_of_mem_bind +#align stream.wseq.exists_of_mem_bind Stream'.Wseq.exists_of_mem_bind theorem destruct_map (f : α → β) (s : Wseq α) : destruct (map f s) = Computation.map (Option.map (Prod.map f (map f))) (destruct s) := by apply - eq_of_bisim fun c1 c2 => + Computation.eq_of_bisim fun c1 c2 => ∃ s, c1 = destruct (map f s) ∧ c2 = Computation.map (Option.map (Prod.map f (map f))) (destruct s) @@ -1447,7 +1453,7 @@ theorem destruct_map (f : α → β) (s : Wseq α) : apply s.rec_on _ (fun a s => _) fun s => _ <;> simp exact ⟨s, rfl, rfl⟩ · exact ⟨s, rfl, rfl⟩ -#align wseq.destruct_map Wseq.destruct_map +#align stream.wseq.destruct_map Stream'.Wseq.destruct_map theorem liftRel_map {δ} (R : α → β → Prop) (S : γ → δ → Prop) {s1 : Wseq α} {s2 : Wseq β} {f1 : α → γ} {f2 : β → δ} (h1 : LiftRel R s1 s2) (h2 : ∀ {a b}, R a b → S (f1 a) (f2 b)) : @@ -1464,25 +1470,26 @@ theorem liftRel_map {δ} (R : α → β → Prop) (S : γ → δ → Prop) {s1 : · cases' a with a s <;> cases' b with b t cases' h with r h exact ⟨h2 r, s, rfl, t, rfl, h⟩⟩ -#align wseq.lift_rel_map Wseq.liftRel_map +#align stream.wseq.lift_rel_map Stream'.Wseq.liftRel_map /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem map_congr (f : α → β) {s t : Wseq α} (h : s ~ t) : map f s ~ map f t := liftRel_map _ _ h fun _ _ => congr_arg _ -#align wseq.map_congr Wseq.map_congr +#align stream.wseq.map_congr Stream'.Wseq.map_congr +/-- auxilary defintion of `destruct_append` over weak sequences-/ @[simp] def DestructAppend.aux (t : Wseq α) : Option (α × Wseq α) → Computation (Option (α × Wseq α)) | none => destruct t | some (a, s) => return (some (a, append s t)) -#align wseq.destruct_append.aux Wseq.DestructAppend.aux +#align stream.wseq.destruct_append.aux Stream'.Wseq.DestructAppend.aux theorem destruct_append (s t : Wseq α) : destruct (append s t) = (destruct s).bind (DestructAppend.aux t) := by apply - eq_of_bisim + Computation.eq_of_bisim (fun c1 c2 => ∃ s t, c1 = destruct (append s t) ∧ c2 = (destruct s).bind (destruct_append.aux t)) _ ⟨s, t, rfl, rfl⟩ @@ -1491,19 +1498,20 @@ theorem destruct_append (s t : Wseq α) : · apply t.rec_on _ (fun b t => _) fun t => _ <;> simp · refine' ⟨nil, t, _, _⟩ <;> simp · exact ⟨s, t, rfl, rfl⟩ -#align wseq.destruct_append Wseq.destruct_append +#align stream.wseq.destruct_append Stream'.Wseq.destruct_append +/-- auxilary defintion of `destruct_join` over weak sequences-/ @[simp] def DestructJoin.aux : Option (Wseq α × Wseq (Wseq α)) → Computation (Option (α × Wseq α)) | none => return none | some (s, S) => (destruct (append s (join S))).think -#align wseq.destruct_join.aux Wseq.DestructJoin.aux +#align stream.wseq.destruct_join.aux Stream'.Wseq.DestructJoin.aux theorem destruct_join (S : Wseq (Wseq α)) : destruct (join S) = (destruct S).bind DestructJoin.aux := by apply - eq_of_bisim + Computation.eq_of_bisim (fun c1 c2 => c1 = c2 ∨ ∃ S, c1 = destruct (join S) ∧ c2 = (destruct S).bind destruct_join.aux) _ (Or.inr ⟨S, rfl, rfl⟩) @@ -1515,7 +1523,7 @@ theorem destruct_join (S : Wseq (Wseq α)) : by apply S.rec_on _ (fun s S => _) fun S => _ <;> simp · refine' Or.inr ⟨S, rfl, rfl⟩ -#align wseq.destruct_join Wseq.destruct_join +#align stream.wseq.destruct_join Stream'.Wseq.destruct_join theorem liftRel_append (R : α → β → Prop) {s1 s2 : Wseq α} {t1 t2 : Wseq β} (h1 : LiftRel R s1 t1) (h2 : LiftRel R s2 t2) : LiftRel R (append s1 s2) (append t1 t2) := @@ -1545,7 +1553,7 @@ theorem liftRel_append (R : α → β → Prop) {s1 s2 : Wseq α} {t1 t2 : Wseq cases' h with r h simp exact ⟨r, Or.inr ⟨s, rfl, t, rfl, h⟩⟩⟩ -#align wseq.lift_rel_append Wseq.liftRel_append +#align stream.wseq.lift_rel_append Stream'.Wseq.liftRel_append theorem LiftRelJoin.lem (R : α → β → Prop) {S T} {U : Wseq α → Wseq β → Prop} (ST : LiftRel (LiftRel R) S T) @@ -1600,7 +1608,7 @@ theorem LiftRelJoin.lem (R : α → β → Prop) {S T} {U : Wseq α → Wseq β apply ret_mem rw [eq_of_ret_mem rs5.mem] exact ⟨ab, HU _ _ ⟨s', t', S', T', rfl, rfl, st', ST'⟩⟩ -#align wseq.lift_rel_join.lem Wseq.LiftRelJoin.lem +#align stream.wseq.lift_rel_join.lem Stream'.Wseq.LiftRelJoin.lem theorem liftRel_join (R : α → β → Prop) {S : Wseq (Wseq α)} {T : Wseq (Wseq β)} (h : LiftRel (LiftRel R) S T) : LiftRel R (join S) (join T) := @@ -1627,18 +1635,18 @@ theorem liftRel_join (R : α → β → Prop) {S : Wseq (Wseq α)} {T : Wseq (Ws · rw [← lift_rel.swap R, ← lift_rel.swap (lift_rel R)] exact fun s1 s2 ⟨s, t, S, T, h1, h2, st, ST⟩ => ⟨t, s, T, S, h2, h1, st, ST⟩ · exact mb⟩ -#align wseq.lift_rel_join Wseq.liftRel_join +#align stream.wseq.lift_rel_join Stream'.Wseq.liftRel_join /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem join_congr {S T : Wseq (Wseq α)} (h : LiftRel Equiv S T) : join S ~ join T := liftRel_join _ h -#align wseq.join_congr Wseq.join_congr +#align stream.wseq.join_congr Stream'.Wseq.join_congr theorem liftRel_bind {δ} (R : α → β → Prop) (S : γ → δ → Prop) {s1 : Wseq α} {s2 : Wseq β} {f1 : α → Wseq γ} {f2 : β → Wseq δ} (h1 : LiftRel R s1 s2) (h2 : ∀ {a b}, R a b → LiftRel S (f1 a) (f2 b)) : LiftRel S (bind s1 f1) (bind s2 f2) := liftRel_join _ (liftRel_map _ _ h1 @h2) -#align wseq.lift_rel_bind Wseq.liftRel_bind +#align stream.wseq.lift_rel_bind Stream'.Wseq.liftRel_bind /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @@ -1646,12 +1654,12 @@ theorem liftRel_bind {δ} (R : α → β → Prop) (S : γ → δ → Prop) {s1 theorem bind_congr {s1 s2 : Wseq α} {f1 f2 : α → Wseq β} (h1 : s1 ~ s2) (h2 : ∀ a, f1 a ~ f2 a) : bind s1 f1 ~ bind s2 f2 := liftRel_bind _ _ h1 fun a b h => by rw [h] <;> apply h2 -#align wseq.bind_congr Wseq.bind_congr +#align stream.wseq.bind_congr Stream'.Wseq.bind_congr /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] theorem join_ret (s : Wseq α) : join (ret s) ~ s := by simp [ret] <;> apply think_equiv -#align wseq.join_ret Wseq.join_ret +#align stream.wseq.join_ret Stream'.Wseq.join_ret /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] @@ -1667,12 +1675,12 @@ theorem join_map_ret (s : Wseq α) : join (map ret s) ~ s := clear h _match have : ∀ s, - ∃ s' : Wseq α, + ∃ s' : wseq α, (map ret s).join.destruct = (map ret s').join.destruct ∧ destruct s = s'.destruct := fun s => ⟨s, rfl, rfl⟩ apply s.rec_on _ (fun a s => _) fun s => _ <;> simp [ret, ret_mem, this, Option.exists] · exact ⟨s, rfl, rfl⟩ -#align wseq.join_map_ret Wseq.join_map_ret +#align stream.wseq.join_map_ret Stream'.Wseq.join_map_ret /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] @@ -1686,7 +1694,7 @@ theorem join_append (S T : Wseq (Wseq α)) : join (append S T) ~ append (join S) apply lift_rel_rec (fun c1 c2 => - ∃ (s : Wseq α)(S T : _), + ∃ (s : wseq α)(S T : _), c1 = destruct (append s (join (append S T))) ∧ c2 = destruct (append s (append (join S) (join T)))) _ _ _ @@ -1697,16 +1705,16 @@ theorem join_append (S T : Wseq (Wseq α)) : join (append S T) ~ append (join S) match c1, c2, h with | _, _, ⟨s, S, T, rfl, rfl⟩ => by clear _match h h - apply Wseq.recOn s _ (fun a s => _) fun s => _ <;> simp - · apply Wseq.recOn S _ (fun s S => _) fun S => _ <;> simp - · apply Wseq.recOn T _ (fun s T => _) fun T => _ <;> simp + apply wseq.rec_on s _ (fun a s => _) fun s => _ <;> simp + · apply wseq.rec_on S _ (fun s S => _) fun S => _ <;> simp + · apply wseq.rec_on T _ (fun s T => _) fun T => _ <;> simp · refine' ⟨s, nil, T, _, _⟩ <;> simp · refine' ⟨nil, nil, T, _, _⟩ <;> simp · exact ⟨s, S, T, rfl, rfl⟩ · refine' ⟨nil, S, T, _, _⟩ <;> simp · exact ⟨s, S, T, rfl, rfl⟩ · exact ⟨s, S, T, rfl, rfl⟩ -#align wseq.join_append Wseq.join_append +#align stream.wseq.join_append Stream'.Wseq.join_append /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] @@ -1714,32 +1722,32 @@ theorem bind_ret (f : α → β) (s) : bind s (ret ∘ f) ~ map f s := by dsimp [bind]; change fun x => ret (f x) with ret ∘ f rw [map_comp]; apply join_map_ret -#align wseq.bind_ret Wseq.bind_ret +#align stream.wseq.bind_ret Stream'.Wseq.bind_ret /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] theorem ret_bind (a : α) (f : α → Wseq β) : bind (ret a) f ~ f a := by simp [bind] -#align wseq.ret_bind Wseq.ret_bind +#align stream.wseq.ret_bind Stream'.Wseq.ret_bind @[simp] theorem map_join (f : α → β) (S) : map f (join S) = join (map (map f) S) := by apply - SeqCat.eq_of_bisim fun s1 s2 => + seq.eq_of_bisim fun s1 s2 => ∃ s S, s1 = append s (map f (join S)) ∧ s2 = append s (join (map (map f) S)) · intro s1 s2 h exact match s1, s2, h with | _, _, ⟨s, S, rfl, rfl⟩ => by - apply Wseq.recOn s _ (fun a s => _) fun s => _ <;> simp - · apply Wseq.recOn S _ (fun s S => _) fun S => _ <;> simp + apply wseq.rec_on s _ (fun a s => _) fun s => _ <;> simp + · apply wseq.rec_on S _ (fun s S => _) fun S => _ <;> simp · exact ⟨map f s, S, rfl, rfl⟩ · refine' ⟨nil, S, _, _⟩ <;> simp · exact ⟨_, _, rfl, rfl⟩ · exact ⟨_, _, rfl, rfl⟩ · refine' ⟨nil, S, _, _⟩ <;> simp -#align wseq.map_join Wseq.map_join +#align stream.wseq.map_join Stream'.Wseq.map_join /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] @@ -1766,16 +1774,16 @@ theorem join_join (SS : Wseq (Wseq (Wseq α))) : join (join SS) ~ join (map join match c1, c2, h with | _, _, ⟨s, S, SS, rfl, rfl⟩ => by clear _match h h - apply Wseq.recOn s _ (fun a s => _) fun s => _ <;> simp - · apply Wseq.recOn S _ (fun s S => _) fun S => _ <;> simp - · apply Wseq.recOn SS _ (fun S SS => _) fun SS => _ <;> simp + apply wseq.rec_on s _ (fun a s => _) fun s => _ <;> simp + · apply wseq.rec_on S _ (fun s S => _) fun S => _ <;> simp + · apply wseq.rec_on SS _ (fun S SS => _) fun SS => _ <;> simp · refine' ⟨nil, S, SS, _, _⟩ <;> simp · refine' ⟨nil, nil, SS, _, _⟩ <;> simp · exact ⟨s, S, SS, rfl, rfl⟩ · refine' ⟨nil, S, SS, _, _⟩ <;> simp · exact ⟨s, S, SS, rfl, rfl⟩ · exact ⟨s, S, SS, rfl, rfl⟩ -#align wseq.join_join Wseq.join_join +#align stream.wseq.join_join Stream'.Wseq.join_join /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[simp] @@ -1784,7 +1792,7 @@ theorem bind_assoc (s : Wseq α) (f : α → Wseq β) (g : β → Wseq γ) : by simp [bind]; rw [← map_comp f (map g), map_comp (map g ∘ f) join] apply join_join -#align wseq.bind_assoc Wseq.bind_assoc +#align stream.wseq.bind_assoc Stream'.Wseq.bind_assoc instance : Monad Wseq where map := @map @@ -1807,3 +1815,5 @@ instance : is_lawful_monad wseq := -/ end Wseq +end Stream' + diff --git a/Mathbin/Data/Set/Pointwise/Smul.lean b/Mathbin/Data/Set/Pointwise/Smul.lean index fc87dd25cc..641d933426 100644 --- a/Mathbin/Data/Set/Pointwise/Smul.lean +++ b/Mathbin/Data/Set/Pointwise/Smul.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Floris van Doorn ! This file was ported from Lean 3 source module data.set.pointwise.smul -! leanprover-community/mathlib commit c227d107bbada5d0d9d20287e3282c0a7f1651a0 +! leanprover-community/mathlib commit b685f506164f8d17a6404048bc4d696739c5d976 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -703,10 +703,12 @@ instance isScalarTower'' [SMul α β] [SMul α γ] [SMul β γ] [IsScalarTower -/ #print Set.isCentralScalar /- +@[to_additive] instance isCentralScalar [SMul α β] [SMul αᵐᵒᵖ β] [IsCentralScalar α β] : IsCentralScalar α (Set β) := ⟨fun a S => (congr_arg fun f => f '' S) <| funext fun _ => op_smul_eq_smul _ _⟩ #align set.is_central_scalar Set.isCentralScalar +#align set.is_central_vadd Set.is_central_vadd -/ #print Set.mulAction /- diff --git a/Mathbin/Geometry/Euclidean/Basic.lean b/Mathbin/Geometry/Euclidean/Basic.lean index d1ebba1c0f..f864f6025c 100644 --- a/Mathbin/Geometry/Euclidean/Basic.lean +++ b/Mathbin/Geometry/Euclidean/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Manuel Candales ! This file was ported from Lean 3 source module geometry.euclidean.basic -! leanprover-community/mathlib commit 46b633fd842bef9469441c0209906f6dddd2b4f5 +! leanprover-community/mathlib commit 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -101,7 +101,7 @@ in terms of the pairwise distances between the points in that combination. -/ theorem dist_affineCombination {ι : Type _} {s : Finset ι} {w₁ w₂ : ι → ℝ} (p : ι → P) (h₁ : (∑ i in s, w₁ i) = 1) (h₂ : (∑ i in s, w₂ i) = 1) : by - have a₁ := s.affine_combination p w₁ <;> have a₂ := s.affine_combination p w₂ <;> + have a₁ := s.affine_combination ℝ p w₁ <;> have a₂ := s.affine_combination ℝ p w₂ <;> exact dist a₁ a₂ * dist a₁ a₂ = (-∑ i₁ in s, @@ -109,7 +109,7 @@ theorem dist_affineCombination {ι : Type _} {s : Finset ι} {w₁ w₂ : ι → (w₁ - w₂) i₁ * (w₁ - w₂) i₂ * (dist (p i₁) (p i₂) * dist (p i₁) (p i₂))) / 2 := by - rw [dist_eq_norm_vsub V (s.affine_combination p w₁) (s.affine_combination p w₂), ← + rw [dist_eq_norm_vsub V (s.affine_combination ℝ p w₁) (s.affine_combination ℝ p w₂), ← @inner_self_eq_norm_mul_norm ℝ, Finset.affineCombination_vsub] have h : (∑ i in s, (w₁ - w₂) i) = 0 := by simp_rw [Pi.sub_apply, Finset.sum_sub_distrib, h₁, h₂, sub_self] diff --git a/Mathbin/Geometry/Euclidean/Circumcenter.lean b/Mathbin/Geometry/Euclidean/Circumcenter.lean index 610515192c..f98703869f 100644 --- a/Mathbin/Geometry/Euclidean/Circumcenter.lean +++ b/Mathbin/Geometry/Euclidean/Circumcenter.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers ! This file was ported from Lean 3 source module geometry.euclidean.circumcenter -! leanprover-community/mathlib commit 46b633fd842bef9469441c0209906f6dddd2b4f5 +! leanprover-community/mathlib commit 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -620,7 +620,7 @@ include V theorem point_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} (s : Simplex ℝ P n) (i : Fin (n + 1)) : s.points i = - (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination s.pointsWithCircumcenter + (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination ℝ s.pointsWithCircumcenter (pointWeightsWithCircumcenter i) := by rw [← points_with_circumcenter_point] @@ -663,7 +663,7 @@ include V theorem centroid_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} (s : Simplex ℝ P n) (fs : Finset (Fin (n + 1))) : fs.centroid ℝ s.points = - (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination s.pointsWithCircumcenter + (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination ℝ s.pointsWithCircumcenter (centroidWeightsWithCircumcenter fs) := by simp_rw [centroid_def, affine_combination_apply, weighted_vsub_of_point_apply, @@ -701,7 +701,7 @@ include V `points_with_circumcenter`. -/ theorem circumcenter_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} (s : Simplex ℝ P n) : s.circumcenter = - (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination s.pointsWithCircumcenter + (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination ℝ s.pointsWithCircumcenter (circumcenterWeightsWithCircumcenter n) := by rw [← points_with_circumcenter_eq_circumcenter] @@ -739,7 +739,7 @@ terms of `points_with_circumcenter`. -/ theorem reflection_circumcenter_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} (s : Simplex ℝ P n) {i₁ i₂ : Fin (n + 1)} (h : i₁ ≠ i₂) : reflection (affineSpan ℝ (s.points '' {i₁, i₂})) s.circumcenter = - (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination s.pointsWithCircumcenter + (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination ℝ s.pointsWithCircumcenter (reflectionCircumcenterWeightsWithCircumcenter i₁ i₂) := by have hc : card ({i₁, i₂} : Finset (Fin (n + 1))) = 2 := by simp [h] diff --git a/Mathbin/Geometry/Euclidean/MongePoint.lean b/Mathbin/Geometry/Euclidean/MongePoint.lean index 27489955ec..b70375a3c7 100644 --- a/Mathbin/Geometry/Euclidean/MongePoint.lean +++ b/Mathbin/Geometry/Euclidean/MongePoint.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers ! This file was ported from Lean 3 source module geometry.euclidean.monge_point -! leanprover-community/mathlib commit 46b633fd842bef9469441c0209906f6dddd2b4f5 +! leanprover-community/mathlib commit 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -141,7 +141,7 @@ include V theorem mongePoint_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} (s : Simplex ℝ P (n + 2)) : s.mongePoint = - (univ : Finset (PointsWithCircumcenterIndex (n + 2))).affineCombination + (univ : Finset (PointsWithCircumcenterIndex (n + 2))).affineCombination ℝ s.pointsWithCircumcenter (mongePointWeightsWithCircumcenter n) := by rw [monge_point_eq_smul_vsub_vadd_circumcenter, diff --git a/Mathbin/Geometry/Manifold/ContMdiff.lean b/Mathbin/Geometry/Manifold/ContMdiff.lean index c2e12b8d46..37dba8aadb 100644 --- a/Mathbin/Geometry/Manifold/ContMdiff.lean +++ b/Mathbin/Geometry/Manifold/ContMdiff.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn ! This file was ported from Lean 3 source module geometry.manifold.cont_mdiff -! leanprover-community/mathlib commit be2c24f56783935652cefffb4bfca7e4b25d167e +! leanprover-community/mathlib commit 0187644979f2d3e10a06e916a869c994facd9a87 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -1534,6 +1534,12 @@ theorem contMdiffWithinAt_fst {s : Set (M × N)} {p : M × N} : · simp only [mfld_simps] #align cont_mdiff_within_at_fst contMdiffWithinAt_fst +theorem ContMdiffWithinAt.fst {f : N → M × M'} {s : Set N} {x : N} + (hf : ContMdiffWithinAt J (I.Prod I') n f s x) : + ContMdiffWithinAt J I n (fun x => (f x).1) s x := + contMdiffWithinAt_fst.comp x hf (mapsTo_image f s) +#align cont_mdiff_within_at.fst ContMdiffWithinAt.fst + theorem contMdiffAt_fst {p : M × N} : ContMdiffAt (I.Prod J) I n Prod.fst p := contMdiffWithinAt_fst #align cont_mdiff_at_fst contMdiffAt_fst @@ -1592,6 +1598,12 @@ theorem contMdiffWithinAt_snd {s : Set (M × N)} {p : M × N} : · simp only [mfld_simps] #align cont_mdiff_within_at_snd contMdiffWithinAt_snd +theorem ContMdiffWithinAt.snd {f : N → M × M'} {s : Set N} {x : N} + (hf : ContMdiffWithinAt J (I.Prod I') n f s x) : + ContMdiffWithinAt J I' n (fun x => (f x).2) s x := + contMdiffWithinAt_snd.comp x hf (mapsTo_image f s) +#align cont_mdiff_within_at.snd ContMdiffWithinAt.snd + theorem contMdiffAt_snd {p : M × N} : ContMdiffAt (I.Prod J) J n Prod.snd p := contMdiffWithinAt_snd #align cont_mdiff_at_snd contMdiffAt_snd @@ -1656,6 +1668,22 @@ theorem smooth_prod_assoc : end Projections +theorem contMdiffWithinAt_prod_iff (f : M → M' × N') {s : Set M} {x : M} : + ContMdiffWithinAt I (I'.Prod J') n f s x ↔ + ContMdiffWithinAt I I' n (Prod.fst ∘ f) s x ∧ ContMdiffWithinAt I J' n (Prod.snd ∘ f) s x := + by + refine' ⟨fun h => ⟨h.fst, h.snd⟩, fun h => _⟩ + simpa only [Prod.mk.eta] using h.1.prod_mk h.2 +#align cont_mdiff_within_at_prod_iff contMdiffWithinAt_prod_iff + +theorem contMdiffAt_prod_iff (f : M → M' × N') {x : M} : + ContMdiffAt I (I'.Prod J') n f x ↔ + ContMdiffAt I I' n (Prod.fst ∘ f) x ∧ ContMdiffAt I J' n (Prod.snd ∘ f) x := + by + simp_rw [← contMdiffWithinAt_univ] + exact contMdiffWithinAt_prod_iff f +#align cont_mdiff_at_prod_iff contMdiffAt_prod_iff + section Prod_map variable {g : N → N'} {r : Set N} {y : N} diff --git a/Mathbin/Geometry/Manifold/ContMdiffMfderiv.lean b/Mathbin/Geometry/Manifold/ContMdiffMfderiv.lean index 3af2d2e3ef..9d99a8f1cc 100644 --- a/Mathbin/Geometry/Manifold/ContMdiffMfderiv.lean +++ b/Mathbin/Geometry/Manifold/ContMdiffMfderiv.lean @@ -1,14 +1,13 @@ /- Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Floris van Doorn ! This file was ported from Lean 3 source module geometry.manifold.cont_mdiff_mfderiv -! leanprover-community/mathlib commit f2ce6086713c78a7f880485f7917ea547a215982 +! leanprover-community/mathlib commit 0187644979f2d3e10a06e916a869c994facd9a87 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.Geometry.Manifold.ContMdiff import Mathbin.Geometry.Manifold.Mfderiv /-! @@ -26,9 +25,9 @@ and related notions. -/ -open Set Function Filter ChartedSpace SmoothManifoldWithCorners +open Set Function Filter ChartedSpace SmoothManifoldWithCorners Bundle -open Topology Manifold +open Topology Manifold Bundle /-! ### Definition of smooth functions between manifolds -/ @@ -54,6 +53,10 @@ variable {𝕜 : Type _} [NontriviallyNormedField 𝕜] [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] {G' : Type _} [TopologicalSpace G'] {J' : ModelWithCorners 𝕜 F' G'} {N' : Type _} [TopologicalSpace N'] [ChartedSpace G' N'] [J's : SmoothManifoldWithCorners J' N'] + -- declare some additional normed spaces, used for fibers of vector bundles + {F₁ : Type _} + [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] {F₂ : Type _} [NormedAddCommGroup F₂] + [NormedSpace 𝕜 F₂] -- declare functions, sets, points and smoothness indices {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} @@ -121,7 +124,7 @@ space are model spaces in models with corners. The general fact is proved in `cont_mdiff_on.continuous_on_tangent_map_within`-/ theorem ContMdiffOn.continuousOn_tangentMapWithin_aux {f : H → H'} {s : Set H} (hf : ContMdiffOn I I' n f s) (hn : 1 ≤ n) (hs : UniqueMdiffOn I s) : - ContinuousOn (tangentMapWithin I I' f s) (TangentBundle.proj I H ⁻¹' s) := + ContinuousOn (tangentMapWithin I I' f s) (π (TangentSpace I) ⁻¹' s) := by suffices h : ContinuousOn @@ -138,7 +141,7 @@ theorem ContMdiffOn.continuousOn_tangentMapWithin_aux {f : H → H'} {s : Set H} ((tangentBundleModelSpaceHomeomorph H' I').symm.Continuous.comp_continuousOn h).comp' A have : univ ∩ ⇑(tangentBundleModelSpaceHomeomorph H I) ⁻¹' (Prod.fst ⁻¹' s) = - TangentBundle.proj I H ⁻¹' s := + π (TangentSpace I) ⁻¹' s := by ext ⟨x, v⟩ simp only [mfld_simps] @@ -200,7 +203,7 @@ are model spaces in models with corners. The general fact is proved in `cont_mdiff_on.cont_mdiff_on_tangent_map_within` -/ theorem ContMdiffOn.contMdiffOn_tangentMapWithin_aux {f : H → H'} {s : Set H} (hf : ContMdiffOn I I' n f s) (hmn : m + 1 ≤ n) (hs : UniqueMdiffOn I s) : - ContMdiffOn I.tangent I'.tangent m (tangentMapWithin I I' f s) (TangentBundle.proj I H ⁻¹' s) := + ContMdiffOn I.tangent I'.tangent m (tangentMapWithin I I' f s) (π (TangentSpace I) ⁻¹' s) := by have m_le_n : m ≤ n := by apply le_trans _ hmn @@ -220,7 +223,7 @@ theorem ContMdiffOn.contMdiffOn_tangentMapWithin_aux {f : H → H'} {s : Set H} have A : range I ×ˢ univ ∩ ((Equiv.sigmaEquivProd H E).symm ∘ fun p : E × E => (I.symm p.fst, p.snd)) ⁻¹' - (TangentBundle.proj I H ⁻¹' s) = + (π (TangentSpace I) ⁻¹' s) = (range I ∩ I.symm ⁻¹' s) ×ˢ univ := by ext ⟨x, v⟩ @@ -269,7 +272,7 @@ include Is I's is `C^m` when `m+1 ≤ n`. -/ theorem ContMdiffOn.contMdiffOn_tangentMapWithin (hf : ContMdiffOn I I' n f s) (hmn : m + 1 ≤ n) (hs : UniqueMdiffOn I s) : - ContMdiffOn I.tangent I'.tangent m (tangentMapWithin I I' f s) (TangentBundle.proj I M ⁻¹' s) := + ContMdiffOn I.tangent I'.tangent m (tangentMapWithin I I' f s) (π (TangentSpace I) ⁻¹' s) := by /- The strategy of the proof is to avoid unfolding the definitions, and reduce by functoriality to the case of functions on the model spaces, where we have already proved the result. @@ -299,28 +302,28 @@ theorem ContMdiffOn.contMdiffOn_tangentMapWithin (hf : ContMdiffOn I I' n f s) ( -- First step: local reduction on the space, to a set `s'` which is contained in chart domains. refine' contMdiffOn_of_locally_contMdiffOn fun p hp => _ have hf' := contMdiffOn_iff.1 hf - simp [TangentBundle.proj] at hp - let l := chart_at H p.1 + simp only [mfld_simps] at hp + let l := chart_at H p.proj set Dl := chart_at (ModelProd H E) p with hDl - let r := chart_at H' (f p.1) + let r := chart_at H' (f p.proj) let Dr := chart_at (ModelProd H' E') (tangentMapWithin I I' f s p) let il := chart_at (ModelProd H E) (tangentMap I I l p) let ir := chart_at (ModelProd H' E') (tangentMap I I' (r ∘ f) p) let s' := f ⁻¹' r.source ∩ s ∩ l.source - let s'_lift := TangentBundle.proj I M ⁻¹' s' + let s'_lift := π (TangentSpace I) ⁻¹' s' let s'l := l.target ∩ l.symm ⁻¹' s' - let s'l_lift := TangentBundle.proj I H ⁻¹' s'l + let s'l_lift := π (TangentSpace I) ⁻¹' s'l rcases continuousOn_iff'.1 hf'.1 r.source r.open_source with ⟨o, o_open, ho⟩ suffices h : ContMdiffOn I.tangent I'.tangent m (tangentMapWithin I I' f s) s'_lift - · refine' ⟨TangentBundle.proj I M ⁻¹' (o ∩ l.source), _, _, _⟩ - show IsOpen (TangentBundle.proj I M ⁻¹' (o ∩ l.source)) - exact (IsOpen.inter o_open l.open_source).Preimage (tangentBundle_proj_continuous _ _) - show p ∈ TangentBundle.proj I M ⁻¹' (o ∩ l.source) - · simp [TangentBundle.proj] - have : p.1 ∈ f ⁻¹' r.source ∩ s := by simp [hp] + · refine' ⟨π (TangentSpace I) ⁻¹' (o ∩ l.source), _, _, _⟩ + show IsOpen (π (TangentSpace I) ⁻¹' (o ∩ l.source)) + exact (IsOpen.inter o_open l.open_source).Preimage (continuous_proj E _) + show p ∈ π (TangentSpace I) ⁻¹' (o ∩ l.source) + · simp + have : p.proj ∈ f ⁻¹' r.source ∩ s := by simp [hp] rw [ho] at this exact this.1 - · have : TangentBundle.proj I M ⁻¹' s ∩ TangentBundle.proj I M ⁻¹' (o ∩ l.source) = s'_lift := + · have : π (TangentSpace I) ⁻¹' s ∩ π (TangentSpace I) ⁻¹' (o ∩ l.source) = s'_lift := by dsimp only [s'_lift, s'] rw [ho] @@ -364,8 +367,8 @@ theorem ContMdiffOn.contMdiffOn_tangentMapWithin (hf : ContMdiffOn I I' n f s) ( by have A : ContMdiffOn I'.tangent I'.tangent m Dr.symm Dr.target := contMdiffOn_chart_symm apply ContMdiffOn.comp A diff_irrfl_lift fun p hp => _ - simp only [s'l_lift, TangentBundle.proj, mfld_simps] at hp - simp only [ir, @LocalEquiv.refl_coe (ModelProd H' E'), hp, mfld_simps] + simp only [s'l_lift, mfld_simps] at hp + simp only [ir, hp, mfld_simps] -- conclusion of this step: the composition of all the maps above is smooth have diff_DrirrflilDl : ContMdiffOn I.tangent I'.tangent m @@ -375,14 +378,14 @@ theorem ContMdiffOn.contMdiffOn_tangentMapWithin (hf : ContMdiffOn I I' n f s) ( have A' : ContMdiffOn I.tangent I.tangent m Dl s'_lift := by apply A.mono fun p hp => _ - simp only [s'_lift, TangentBundle.proj, mfld_simps] at hp + simp only [s'_lift, mfld_simps] at hp simp only [Dl, hp, mfld_simps] have B : ContMdiffOn I.tangent I.tangent m il.symm il.target := contMdiffOn_chart_symm have C : ContMdiffOn I.tangent I.tangent m (il.symm ∘ Dl) s'_lift := ContMdiffOn.comp B A' fun p hp => by simp only [il, mfld_simps] apply ContMdiffOn.comp diff_Drirrfl_lift C fun p hp => _ - simp only [s'_lift, TangentBundle.proj, mfld_simps] at hp - simp only [il, s'l_lift, hp, TangentBundle.proj, mfld_simps] + simp only [s'_lift, mfld_simps] at hp + simp only [il, s'l_lift, hp, total_space.proj, mfld_simps] /- Third step: check that the composition of all the maps indeed coincides with the derivative we are looking for -/ have eq_comp : @@ -391,7 +394,7 @@ theorem ContMdiffOn.contMdiffOn_tangentMapWithin (hf : ContMdiffOn I I' n f s) ( (Dr.symm ∘ ir ∘ tangentMapWithin I I' (r ∘ f ∘ l.symm) s'l ∘ il.symm ∘ Dl) q := by intro q hq - simp only [s'_lift, TangentBundle.proj, mfld_simps] at hq + simp only [s'_lift, mfld_simps] at hq have U'q : UniqueMdiffWithinAt I s' q.1 := by apply U' simp only [hq, s', mfld_simps] @@ -447,7 +450,7 @@ theorem ContMdiffOn.contMdiffOn_tangentMapWithin (hf : ContMdiffOn I I' n f s) ( simp [hq] · refine' mdifferentiableAt_atlas _ (chart_mem_atlas _ _) _ simp only [hq, mfld_simps] - have : f p.1 = (tangentMapWithin I I' f s p).1 := rfl + have : f p.proj = (tangentMapWithin I I' f s p).1 := rfl rw [A] dsimp [r, Dr] rw [this, tangentMap_chart] @@ -471,9 +474,9 @@ theorem ContMdiffOn.contMdiffOn_tangentMapWithin (hf : ContMdiffOn I I' n f s) ( derivative is continuous there. -/ theorem ContMdiffOn.continuousOn_tangentMapWithin (hf : ContMdiffOn I I' n f s) (hmn : 1 ≤ n) (hs : UniqueMdiffOn I s) : - ContinuousOn (tangentMapWithin I I' f s) (TangentBundle.proj I M ⁻¹' s) := + ContinuousOn (tangentMapWithin I I' f s) (π (TangentSpace I) ⁻¹' s) := haveI : - ContMdiffOn I.tangent I'.tangent 0 (tangentMapWithin I I' f s) (TangentBundle.proj I M ⁻¹' s) := + ContMdiffOn I.tangent I'.tangent 0 (tangentMapWithin I I' f s) (π (TangentSpace I) ⁻¹' s) := hf.cont_mdiff_on_tangent_map_within hmn hs this.continuous_on #align cont_mdiff_on.continuous_on_tangent_map_within ContMdiffOn.continuousOn_tangentMapWithin @@ -499,172 +502,12 @@ theorem ContMdiff.continuous_tangentMap (hf : ContMdiff I I' n f) (hmn : 1 ≤ n end tangentMap -/-! ### Smoothness of the projection in a basic smooth bundle -/ - - -namespace BasicSmoothVectorBundleCore - -variable (Z : BasicSmoothVectorBundleCore I M E') - -/-- A version of `cont_mdiff_at_iff_target` when the codomain is the total space of - a `basic_smooth_vector_bundle_core`. The continuity condition in the RHS is weaker. -/ -theorem contMdiffAt_iff_target {f : N → Z.toVectorBundleCore.TotalSpace} {x : N} {n : ℕ∞} : - ContMdiffAt J (I.Prod 𝓘(𝕜, E')) n f x ↔ - ContinuousAt (Bundle.TotalSpace.proj ∘ f) x ∧ - ContMdiffAt J 𝓘(𝕜, E × E') n (extChartAt (I.Prod 𝓘(𝕜, E')) (f x) ∘ f) x := - by - let Z' := Z.to_vector_bundle_core - rw [contMdiffAt_iff_target, and_congr_left_iff] - refine' fun hf => ⟨fun h => Z'.continuous_proj.continuous_at.comp h, fun h => _⟩ - exact - (Z'.local_triv ⟨chart_at _ (f x).1, chart_mem_atlas _ _⟩).continuousAt_of_comp_left h - (mem_chart_source _ _) (h.prod hf.continuous_at.snd) -#align basic_smooth_vector_bundle_core.cont_mdiff_at_iff_target BasicSmoothVectorBundleCore.contMdiffAt_iff_target - -theorem smooth_iff_target {f : N → Z.toVectorBundleCore.TotalSpace} : - Smooth J (I.Prod 𝓘(𝕜, E')) f ↔ - Continuous (Bundle.TotalSpace.proj ∘ f) ∧ - ∀ x, SmoothAt J 𝓘(𝕜, E × E') (extChartAt (I.Prod 𝓘(𝕜, E')) (f x) ∘ f) x := - by - simp_rw [Smooth, SmoothAt, ContMdiff, Z.cont_mdiff_at_iff_target, forall_and, - continuous_iff_continuousAt] -#align basic_smooth_vector_bundle_core.smooth_iff_target BasicSmoothVectorBundleCore.smooth_iff_target - -theorem contMdiff_proj : ContMdiff (I.Prod 𝓘(𝕜, E')) I n Z.toVectorBundleCore.proj := - by - intro x - rw [ContMdiffAt, contMdiffWithinAt_iff'] - refine' ⟨Z.to_vector_bundle_core.continuous_proj.continuous_within_at, _⟩ - simp only [(· ∘ ·), chart_at, chart, mfld_simps] - apply cont_diff_within_at_fst.congr - · rintro ⟨a, b⟩ hab - simp only [mfld_simps] at hab - simp only [hab, mfld_simps] - · simp only [mfld_simps] -#align basic_smooth_vector_bundle_core.cont_mdiff_proj BasicSmoothVectorBundleCore.contMdiff_proj - -theorem smooth_proj : Smooth (I.Prod 𝓘(𝕜, E')) I Z.toVectorBundleCore.proj := - contMdiff_proj Z -#align basic_smooth_vector_bundle_core.smooth_proj BasicSmoothVectorBundleCore.smooth_proj - -theorem contMdiffOn_proj {s : Set Z.toVectorBundleCore.TotalSpace} : - ContMdiffOn (I.Prod 𝓘(𝕜, E')) I n Z.toVectorBundleCore.proj s := - Z.contMdiff_proj.ContMdiffOn -#align basic_smooth_vector_bundle_core.cont_mdiff_on_proj BasicSmoothVectorBundleCore.contMdiffOn_proj - -theorem smoothOn_proj {s : Set Z.toVectorBundleCore.TotalSpace} : - SmoothOn (I.Prod 𝓘(𝕜, E')) I Z.toVectorBundleCore.proj s := - contMdiffOn_proj Z -#align basic_smooth_vector_bundle_core.smooth_on_proj BasicSmoothVectorBundleCore.smoothOn_proj - -theorem contMdiffAt_proj {p : Z.toVectorBundleCore.TotalSpace} : - ContMdiffAt (I.Prod 𝓘(𝕜, E')) I n Z.toVectorBundleCore.proj p := - Z.contMdiff_proj.ContMdiffAt -#align basic_smooth_vector_bundle_core.cont_mdiff_at_proj BasicSmoothVectorBundleCore.contMdiffAt_proj - -theorem smoothAt_proj {p : Z.toVectorBundleCore.TotalSpace} : - SmoothAt (I.Prod 𝓘(𝕜, E')) I Z.toVectorBundleCore.proj p := - Z.contMdiffAt_proj -#align basic_smooth_vector_bundle_core.smooth_at_proj BasicSmoothVectorBundleCore.smoothAt_proj - -theorem contMdiffWithinAt_proj {s : Set Z.toVectorBundleCore.TotalSpace} - {p : Z.toVectorBundleCore.TotalSpace} : - ContMdiffWithinAt (I.Prod 𝓘(𝕜, E')) I n Z.toVectorBundleCore.proj s p := - Z.contMdiffAt_proj.ContMdiffWithinAt -#align basic_smooth_vector_bundle_core.cont_mdiff_within_at_proj BasicSmoothVectorBundleCore.contMdiffWithinAt_proj - -theorem smoothWithinAt_proj {s : Set Z.toVectorBundleCore.TotalSpace} - {p : Z.toVectorBundleCore.TotalSpace} : - SmoothWithinAt (I.Prod 𝓘(𝕜, E')) I Z.toVectorBundleCore.proj s p := - Z.contMdiffWithinAt_proj -#align basic_smooth_vector_bundle_core.smooth_within_at_proj BasicSmoothVectorBundleCore.smoothWithinAt_proj - -/-- If an element of `E'` is invariant under all coordinate changes, then one can define a -corresponding section of the fiber bundle, which is smooth. This applies in particular to the -zero section of a vector bundle. Another example (not yet defined) would be the identity -section of the endomorphism bundle of a vector bundle. -/ -theorem smooth_const_section (v : E') - (h : ∀ i j : atlas H M, ∀ x ∈ i.1.source ∩ j.1.source, Z.coordChange i j (i.1 x) v = v) : - Smooth I (I.Prod 𝓘(𝕜, E')) (show M → Z.toVectorBundleCore.TotalSpace from fun x => ⟨x, v⟩) := - by - intro x - rw [ContMdiffAt, contMdiffWithinAt_iff'] - constructor - · apply Continuous.continuousWithinAt - apply FiberBundleCore.continuous_const_section - intro i j y hy - exact h _ _ _ hy - · have : ContDiff 𝕜 ⊤ fun y : E => (y, v) := cont_diff_id.prod contDiff_const - apply this.cont_diff_within_at.congr - · intro y hy - simp only [mfld_simps] at hy - simp only [chart, hy, chart_at, Prod.mk.inj_iff, to_vector_bundle_core, mfld_simps] - apply h - simp only [hy, Subtype.val_eq_coe, mfld_simps] - · simp only [chart, chart_at, Prod.mk.inj_iff, to_vector_bundle_core, mfld_simps] - apply h - simp only [Subtype.val_eq_coe, mfld_simps] -#align basic_smooth_vector_bundle_core.smooth_const_section BasicSmoothVectorBundleCore.smooth_const_section - -end BasicSmoothVectorBundleCore - -/-! ### Smoothness of the tangent bundle projection -/ - - namespace TangentBundle include Is -theorem contMdiff_proj : ContMdiff I.tangent I n (proj I M) := - BasicSmoothVectorBundleCore.contMdiff_proj _ -#align tangent_bundle.cont_mdiff_proj TangentBundle.contMdiff_proj - -theorem smooth_proj : Smooth I.tangent I (proj I M) := - BasicSmoothVectorBundleCore.smooth_proj _ -#align tangent_bundle.smooth_proj TangentBundle.smooth_proj - -theorem contMdiffOn_proj {s : Set (TangentBundle I M)} : ContMdiffOn I.tangent I n (proj I M) s := - BasicSmoothVectorBundleCore.contMdiffOn_proj _ -#align tangent_bundle.cont_mdiff_on_proj TangentBundle.contMdiffOn_proj - -theorem smoothOn_proj {s : Set (TangentBundle I M)} : SmoothOn I.tangent I (proj I M) s := - BasicSmoothVectorBundleCore.smoothOn_proj _ -#align tangent_bundle.smooth_on_proj TangentBundle.smoothOn_proj - -theorem contMdiffAt_proj {p : TangentBundle I M} : ContMdiffAt I.tangent I n (proj I M) p := - BasicSmoothVectorBundleCore.contMdiffAt_proj _ -#align tangent_bundle.cont_mdiff_at_proj TangentBundle.contMdiffAt_proj - -theorem smoothAt_proj {p : TangentBundle I M} : SmoothAt I.tangent I (proj I M) p := - BasicSmoothVectorBundleCore.smoothAt_proj _ -#align tangent_bundle.smooth_at_proj TangentBundle.smoothAt_proj - -theorem contMdiffWithinAt_proj {s : Set (TangentBundle I M)} {p : TangentBundle I M} : - ContMdiffWithinAt I.tangent I n (proj I M) s p := - BasicSmoothVectorBundleCore.contMdiffWithinAt_proj _ -#align tangent_bundle.cont_mdiff_within_at_proj TangentBundle.contMdiffWithinAt_proj - -theorem smoothWithinAt_proj {s : Set (TangentBundle I M)} {p : TangentBundle I M} : - SmoothWithinAt I.tangent I (proj I M) s p := - BasicSmoothVectorBundleCore.smoothWithinAt_proj _ -#align tangent_bundle.smooth_within_at_proj TangentBundle.smoothWithinAt_proj - variable (I M) -/-- The zero section of the tangent bundle -/ -def zeroSection : M → TangentBundle I M := fun x => ⟨x, 0⟩ -#align tangent_bundle.zero_section TangentBundle.zeroSection - -variable {I M} - -theorem smooth_zeroSection : Smooth I I.tangent (zeroSection I M) := - by - apply BasicSmoothVectorBundleCore.smooth_const_section (tangentBundleCore I M) 0 - intro i j x hx - simp only [tangentBundleCore, ContinuousLinearMap.map_zero, ContinuousLinearMap.coe_coe, - mfld_simps] -#align tangent_bundle.smooth_zero_section TangentBundle.smooth_zeroSection - open Bundle /-- The derivative of the zero section of the tangent bundle maps `⟨x, v⟩` to `⟨⟨x, 0⟩, ⟨v, 0⟩⟩`. @@ -682,7 +525,7 @@ may seem. TODO define splittings of vector bundles; state this result invariantly. -/ theorem tangentMap_tangentBundle_pure (p : TangentBundle I M) : - tangentMap I I.tangent (TangentBundle.zeroSection I M) p = ⟨⟨p.1, 0⟩, ⟨p.2, 0⟩⟩ := + tangentMap I I.tangent (zeroSection (TangentSpace I)) p = ⟨⟨p.proj, 0⟩, ⟨p.2, 0⟩⟩ := by rcases p with ⟨x, v⟩ have N : I.symm ⁻¹' (chart_at H x).target ∈ 𝓝 (I ((chart_at H x) x)) := @@ -691,10 +534,11 @@ theorem tangentMap_tangentBundle_pure (p : TangentBundle I M) : apply (LocalHomeomorph.open_target _).Preimage I.continuous_inv_fun simp only [mfld_simps] have A : MdifferentiableAt I I.tangent (fun x => @total_space_mk M (TangentSpace I) x 0) x := - tangent_bundle.smooth_zero_section.mdifferentiable_at + haveI : Smooth I (I.prod 𝓘(𝕜, E)) (zero_section (TangentSpace I : M → Type _)) := + Bundle.smooth_zeroSection 𝕜 (TangentSpace I : M → Type _) + this.mdifferentiable_at have B : - fderivWithin 𝕜 (fun x_1 : E => (x_1, (0 : E))) (Set.range ⇑I) (I ((chart_at H x) x)) v = - (v, 0) := + fderivWithin 𝕜 (fun x' : E => (x', (0 : E))) (Set.range ⇑I) (I ((chart_at H x) x)) v = (v, 0) := by rw [fderivWithin_eq_fderiv, DifferentiableAt.fderiv_prod] · simp @@ -702,9 +546,9 @@ theorem tangentMap_tangentBundle_pure (p : TangentBundle I M) : · exact differentiableAt_const _ · exact ModelWithCorners.unique_diff_at_image I · exact differentiable_at_id'.prod (differentiableAt_const _) - simp only [TangentBundle.zeroSection, tangentMap, mfderiv, A, if_pos, chart_at, - BasicSmoothVectorBundleCore.chart, BasicSmoothVectorBundleCore.toVectorBundleCore, - tangentBundleCore, Function.comp, ContinuousLinearMap.map_zero, mfld_simps] + simp only [Bundle.zeroSection, tangentMap, mfderiv, total_space.proj_mk, A, if_pos, chart_at, + FiberBundle.chartedSpace_chartAt, TangentBundle.trivializationAt_apply, tangentBundleCore, + Function.comp, ContinuousLinearMap.map_zero, mfld_simps] rw [← fderivWithin_inter N (I.unique_diff (I ((chart_at H x) x)) (Set.mem_range_self _))] at B rw [← fderivWithin_inter N (I.unique_diff (I ((chart_at H x) x)) (Set.mem_range_self _)), ← B] congr 2 diff --git a/Mathbin/Geometry/Manifold/Mfderiv.lean b/Mathbin/Geometry/Manifold/Mfderiv.lean index 3b213557b8..355082ffac 100644 --- a/Mathbin/Geometry/Manifold/Mfderiv.lean +++ b/Mathbin/Geometry/Manifold/Mfderiv.lean @@ -1,15 +1,14 @@ /- Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Floris van Doorn ! This file was ported from Lean 3 source module geometry.manifold.mfderiv -! leanprover-community/mathlib commit f2ce6086713c78a7f880485f7917ea547a215982 +! leanprover-community/mathlib commit 0187644979f2d3e10a06e916a869c994facd9a87 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.Geometry.Manifold.LocalInvariantProperties -import Mathbin.Geometry.Manifold.TangentBundle +import Mathbin.Geometry.Manifold.VectorBundle.Tangent /-! # The derivative of functions between smooth manifolds @@ -66,9 +65,9 @@ of `f` in these charts. Due to the fact that we are working in a model with corners, with an additional embedding `I` of the model space `H` in the model vector space `E`, the charts taking values in `E` are not the original charts of the manifold, but those ones composed with `I`, called extended charts. We define -`written_in_ext_chart I I' x f` for the function `f` written in the preferred extended charts. Then -the manifold derivative of `f`, at `x`, is just the usual derivative of `written_in_ext_chart I I' x -f`, at the point `(ext_chart_at I x) x`. +`written_in_ext_chart I I' x f` for the function `f` written in the preferred extended charts. Then +the manifold derivative of `f`, at `x`, is just the usual derivative of +`written_in_ext_chart I I' x f`, at the point `(ext_chart_at I x) x`. There is a subtelty with respect to continuity: if the function is not continuous, then the image of a small open set around `x` will not be contained in the source of the preferred chart around @@ -100,7 +99,7 @@ Derivative, manifold noncomputable section -open Classical Topology Manifold +open Classical Topology Manifold Bundle open Set @@ -721,26 +720,25 @@ theorem tangentMapWithin_eq_tangentMap {p : TangentBundle I M} (hs : UniqueMdiff #align tangent_map_within_eq_tangent_map tangentMapWithin_eq_tangentMap @[simp, mfld_simps] -theorem tangentMapWithin_tangentBundle_proj {p : TangentBundle I M} : - TangentBundle.proj I' M' (tangentMapWithin I I' f s p) = f (TangentBundle.proj I M p) := +theorem tangentMapWithin_proj {p : TangentBundle I M} : + (tangentMapWithin I I' f s p).proj = f p.proj := rfl -#align tangent_map_within_tangent_bundle_proj tangentMapWithin_tangentBundle_proj +#align tangent_map_within_proj tangentMapWithin_proj @[simp, mfld_simps] -theorem tangentMapWithin_proj {p : TangentBundle I M} : (tangentMapWithin I I' f s p).1 = f p.1 := +theorem tangentMapWithin_fst {p : TangentBundle I M} : (tangentMapWithin I I' f s p).1 = f p.1 := rfl -#align tangent_map_within_proj tangentMapWithin_proj +#align tangent_map_within_fst tangentMapWithin_fst @[simp, mfld_simps] -theorem tangentMap_tangentBundle_proj {p : TangentBundle I M} : - TangentBundle.proj I' M' (tangentMap I I' f p) = f (TangentBundle.proj I M p) := +theorem tangentMap_proj {p : TangentBundle I M} : (tangentMap I I' f p).proj = f p.proj := rfl -#align tangent_map_tangent_bundle_proj tangentMap_tangentBundle_proj +#align tangent_map_proj tangentMap_proj @[simp, mfld_simps] -theorem tangentMap_proj {p : TangentBundle I M} : (tangentMap I I' f p).1 = f p.1 := +theorem tangentMap_fst {p : TangentBundle I M} : (tangentMap I I' f p).1 = f p.1 := rfl -#align tangent_map_proj tangentMap_proj +#align tangent_map_fst tangentMap_fst omit Is I's @@ -1273,8 +1271,7 @@ theorem tangentMap_id : tangentMap I I (id : M → M) = id := simp [tangentMap] #align tangent_map_id tangentMap_id -theorem tangentMapWithin_id {p : TangentBundle I M} - (hs : UniqueMdiffWithinAt I s (TangentBundle.proj I M p)) : +theorem tangentMapWithin_id {p : TangentBundle I M} (hs : UniqueMdiffWithinAt I s p.proj) : tangentMapWithin I I (id : M → M) s p = p := by simp only [tangentMapWithin, id.def] @@ -1640,8 +1637,8 @@ theorem tangentMap_chart_symm {p : TangentBundle I M} {q : TangentBundle I H} -- a trivial instance is needed after the rewrite, handle it right now. rotate_left; · infer_instance - simp only [ContinuousLinearMap.coe_coe, BasicSmoothVectorBundleCore.chart, h, tangentBundleCore, - BasicSmoothVectorBundleCore.toVectorBundleCore, chart_at, Sigma.mk.inj_iff, mfld_simps] + simp only [ContinuousLinearMap.coe_coe, TangentBundle.chartAt, h, tangentBundleCore, chart_at, + Sigma.mk.inj_iff, mfld_simps] #align tangent_map_chart_symm tangentMap_chart_symm end Charts @@ -1919,15 +1916,18 @@ theorem UniqueMdiffOn.uniqueDiffOn_inter_preimage (hs : UniqueMdiffOn I s) (x : this.unique_diff_on_target_inter _ #align unique_mdiff_on.unique_diff_on_inter_preimage UniqueMdiffOn.uniqueDiffOn_inter_preimage -variable {F : Type _} [NormedAddCommGroup F] [NormedSpace 𝕜 F] - (Z : BasicSmoothVectorBundleCore I M F) +open Bundle + +variable {F : Type _} [NormedAddCommGroup F] [NormedSpace 𝕜 F] (Z : M → Type _) + [TopologicalSpace (TotalSpace Z)] [∀ b, TopologicalSpace (Z b)] [∀ b, AddCommMonoid (Z b)] + [∀ b, Module 𝕜 (Z b)] [FiberBundle F Z] [VectorBundle 𝕜 F Z] [SmoothVectorBundle F Z I] /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -/-- In a smooth fiber bundle constructed from core, the preimage under the projection of a set with +/-- In a smooth fiber bundle, the preimage under the projection of a set with unique differential in the basis also has unique differential. -/ theorem UniqueMdiffOn.smooth_bundle_preimage (hs : UniqueMdiffOn I s) : - UniqueMdiffOn (I.Prod 𝓘(𝕜, F)) (Z.toVectorBundleCore.proj ⁻¹' s) := + UniqueMdiffOn (I.Prod 𝓘(𝕜, F)) (π Z ⁻¹' s) := by /- Using a chart (and the fact that unique differentiability is invariant under charts), we reduce the situation to the model space, where we can use the fact that products respect @@ -1937,37 +1937,54 @@ theorem UniqueMdiffOn.smooth_bundle_preimage (hs : UniqueMdiffOn I s) : · simpa only [mfld_simps] using hp let e₀ := chart_at H p.1 let e := chart_at (ModelProd H F) p + have h2s : + ∀ x, + x ∈ e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s) ↔ + (x.1 ∈ e₀.target ∧ e₀.symm x.1 ∈ (trivialization_at F Z p.1).baseSet) ∧ e₀.symm x.1 ∈ s := + by + intro x + have A : x ∈ e.target ↔ x.1 ∈ e₀.target ∧ e₀.symm x.1 ∈ (trivialization_at F Z p.1).baseSet := + by + simp only [e, FiberBundle.chartedSpace_chartAt, Trivialization.mem_target, + Bundle.TotalSpace.proj, mfld_simps] + rw [← A, mem_inter_iff, and_congr_right_iff] + intro hx + simp only [FiberBundle.chartedSpace_chartAt_symm_fst p x hx, mfld_simps] -- It suffices to prove unique differentiability in a chart - suffices h : - UniqueMdiffOn (I.prod 𝓘(𝕜, F)) (e.target ∩ e.symm ⁻¹' (Z.to_vector_bundle_core.proj ⁻¹' s)) + suffices h : UniqueMdiffOn (I.prod 𝓘(𝕜, F)) (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s)) · have A : UniqueMdiffOn (I.prod 𝓘(𝕜, F)) - (e.symm.target ∩ - e.symm.symm ⁻¹' (e.target ∩ e.symm ⁻¹' (Z.to_vector_bundle_core.proj ⁻¹' s))) := + (e.symm.target ∩ e.symm.symm ⁻¹' (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s))) := by apply h.unique_mdiff_on_preimage exact (mdifferentiable_of_mem_atlas _ (chart_mem_atlas _ _)).symm infer_instance - have : - p ∈ - e.symm.target ∩ - e.symm.symm ⁻¹' (e.target ∩ e.symm ⁻¹' (Z.to_vector_bundle_core.proj ⁻¹' s)) := - by simp only [e, hp, mfld_simps] + have : p ∈ e.symm.target ∩ e.symm.symm ⁻¹' (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s)) := by + simp only [e, hp, mfld_simps] apply (A _ this).mono intro q hq simp only [e, LocalHomeomorph.left_inv _ hq.1, mfld_simps] at hq simp only [hq, mfld_simps] - -- rewrite the relevant set in the chart as a direct product - have : - (fun p : E × F => (I.symm p.1, p.snd)) ⁻¹' e.target ∩ - (fun p : E × F => (I.symm p.1, p.snd)) ⁻¹' (e.symm ⁻¹' (Sigma.fst ⁻¹' s)) ∩ - range I ×ˢ univ = - (I.symm ⁻¹' (e₀.target ∩ e₀.symm ⁻¹' s) ∩ range I) ×ˢ univ := - by mfld_set_tac intro q hq - replace hq : q.1 ∈ (chart_at H p.1).target ∧ ((chart_at H p.1).symm : H → M) q.1 ∈ s - · simpa only [mfld_simps] using hq - simp only [UniqueMdiffWithinAt, ModelWithCorners.prod, preimage_inter, this, mfld_simps] + simp only [UniqueMdiffWithinAt, ModelWithCorners.prod, -preimage_inter, mfld_simps] + have : + 𝓝[(I.symm ⁻¹' (e₀.target ∩ e₀.symm ⁻¹' s) ∩ range I) ×ˢ univ] (I q.1, q.2) ≤ + 𝓝[(fun p : E × F => (I.symm p.1, p.snd)) ⁻¹' (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s)) ∩ + range I ×ˢ univ] + (I q.1, q.2) := + by + rw [nhdsWithin_le_iff, mem_nhdsWithin] + refine' ⟨(fun p : E × F => (I.symm p.1, p.snd)) ⁻¹' e.target, _, _, _⟩ + · exact e.open_target.preimage (I.continuous_symm.prod_map continuous_id) + · simp only [Prod.mk.eta, mfld_simps] at hq + simp only [Prod.mk.eta, hq, mfld_simps] + rintro x hx + simp only [mfld_simps] at hx + have h2x := hx + simp only [e, FiberBundle.chartedSpace_chartAt, Trivialization.mem_target, mfld_simps] at h2x + simp only [h2s, hx, h2x, -preimage_inter, mfld_simps] + refine' UniqueDiffWithinAt.mono_nhds _ this + rw [h2s] at hq -- apply unique differentiability of products to conclude apply UniqueDiffOn.prod _ uniqueDiffOn_univ · simp only [hq, mfld_simps] @@ -1982,10 +1999,10 @@ theorem UniqueMdiffOn.smooth_bundle_preimage (hs : UniqueMdiffOn I s) : rwa [← preimage_inter, ModelWithCorners.right_inv _ hx.2] at B #align unique_mdiff_on.smooth_bundle_preimage UniqueMdiffOn.smooth_bundle_preimage -theorem UniqueMdiffOn.tangentBundle_proj_preimage (hs : UniqueMdiffOn I s) : - UniqueMdiffOn I.tangent (TangentBundle.proj I M ⁻¹' s) := +theorem UniqueMdiffOn.tangent_bundle_proj_preimage (hs : UniqueMdiffOn I s) : + UniqueMdiffOn I.tangent (π (TangentSpace I) ⁻¹' s) := hs.smooth_bundle_preimage _ -#align unique_mdiff_on.tangent_bundle_proj_preimage UniqueMdiffOn.tangentBundle_proj_preimage +#align unique_mdiff_on.tangent_bundle_proj_preimage UniqueMdiffOn.tangent_bundle_proj_preimage end UniqueMdiff diff --git a/Mathbin/Geometry/Manifold/TangentBundle.lean b/Mathbin/Geometry/Manifold/TangentBundle.lean deleted file mode 100644 index 52a10cb0a7..0000000000 --- a/Mathbin/Geometry/Manifold/TangentBundle.lean +++ /dev/null @@ -1,767 +0,0 @@ -/- -Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel - -! This file was ported from Lean 3 source module geometry.manifold.tangent_bundle -! leanprover-community/mathlib commit f2ce6086713c78a7f880485f7917ea547a215982 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.Topology.VectorBundle.Basic -import Mathbin.Geometry.Manifold.SmoothManifoldWithCorners -import Mathbin.Data.Set.Prod - -/-! -# Basic smooth bundles - -In general, a smooth bundle is a bundle over a smooth manifold, whose fiber is a manifold, and -for which the coordinate changes are smooth. In this definition, there are charts involved at -several places: in the manifold structure of the base, in the manifold structure of the fibers, and -in the local trivializations. This makes it a complicated object in general. There is however a -specific situation where things are much simpler: when the fiber is a vector space (no need for -charts for the fibers), and when the local trivializations of the bundle and the charts of the base -coincide. Then everything is expressed in terms of the charts of the base, making for a much -simpler overall structure, which is easier to manipulate formally. - -Most vector bundles that naturally occur in differential geometry are of this form: -the tangent bundle, the cotangent bundle, differential forms (used to define de Rham cohomology) -and the bundle of Riemannian metrics. Therefore, it is worth defining a specific constructor for -this kind of bundle, that we call basic smooth bundles. - -A basic smooth bundle is thus a smooth bundle over a smooth manifold whose fiber is a vector space, -and which is trivial in the coordinate charts of the base. (We recall that in our notion of manifold -there is a distinguished atlas, which does not need to be maximal: we require the triviality above -this specific atlas). It can be constructed from a basic smooth bundled core, defined below, -specifying the changes in the fiber when one goes from one coordinate chart to another one. - -## Main definitions - -* `basic_smooth_vector_bundle_core I M F`: assuming that `M` is a smooth manifold over the model - with corners `I` on `(𝕜, E, H)`, and `F` is a normed vector space over `𝕜`, this structure - registers, for each pair of charts of `M`, a linear change of coordinates on `F` depending - smoothly on the base point. This is the core structure from which one will build a smooth vector - bundle with fiber `F` over `M`. - -Let `Z` be a basic smooth bundle core over `M` with fiber `F`. We define -`Z.to_vector_bundle_core`, the (topological) vector bundle core associated to `Z`. From -it, we get a space `Z.to_vector_bundle_core.total_space` (which as a Type is just -`Σ (x : M), F`), with the fiber bundle topology. It inherits a manifold structure (where the -charts are in bijection with the charts of the basis). We show that this manifold is smooth. - -Then we use this machinery to construct the tangent bundle of a smooth manifold. - -* `tangent_bundle_core I M`: the basic smooth bundle core associated to a smooth manifold `M` over - a model with corners `I`. -* `tangent_bundle I M` : the total space of `tangent_bundle_core I M`. It is itself a - smooth manifold over the model with corners `I.tangent`, the product of `I` and the trivial model - with corners on `E`. -* `tangent_space I x` : the tangent space to `M` at `x` -* `tangent_bundle.proj I M`: the projection from the tangent bundle to the base manifold - -## Implementation notes - -We register the vector space structure on the fibers of the tangent bundle, but we do not register -the normed space structure coming from that of `F` (as it is not canonical, and we also want to -keep the possibility to add a Riemannian structure on the manifold later on without having two -competing normed space instances on the tangent spaces). - -We require `F` to be a normed space, and not just a topological vector space, as we want to talk -about smooth functions on `F`. The notion of derivative requires a norm to be defined. - -## TODO -construct the cotangent bundle, and the bundles of differential forms. They should follow -functorially from the description of the tangent bundle as a basic smooth bundle. - -## Tags -Smooth fiber bundle, vector bundle, tangent space, tangent bundle --/ - - -noncomputable section - -universe u - -open TopologicalSpace Set - -open Manifold Topology - -/-- Core structure used to create a smooth bundle above `M` (a manifold over the model with -corner `I`) with fiber the normed vector space `F` over `𝕜`, which is trivial in the chart domains -of `M`. This structure registers the changes in the fibers when one changes coordinate charts in the -base. We require the change of coordinates of the fibers to be linear, so that the resulting bundle -is a vector bundle. -/ -structure BasicSmoothVectorBundleCore {𝕜 : Type _} [NontriviallyNormedField 𝕜] {E : Type _} - [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type _} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) (M : Type _) [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] (F : Type _) [NormedAddCommGroup F] [NormedSpace 𝕜 F] where - coordChange : atlas H M → atlas H M → H → F →L[𝕜] F - coordChange_self : ∀ i : atlas H M, ∀ x ∈ i.1.target, ∀ v, coord_change i i x v = v - coordChange_comp : - ∀ i j k : atlas H M, - ∀ x ∈ ((i.1.symm.trans j.1).trans (j.1.symm.trans k.1)).source, - ∀ v, - (coord_change j k ((i.1.symm.trans j.1) x)) (coord_change i j x v) = coord_change i k x v - coordChange_smooth_clm : - ∀ i j : atlas H M, ContDiffOn 𝕜 ∞ (coord_change i j ∘ I.symm) (I '' (i.1.symm.trans j.1).source) -#align basic_smooth_vector_bundle_core BasicSmoothVectorBundleCore - -/-- The trivial basic smooth bundle core, in which all the changes of coordinates are the -identity. -/ -def trivialBasicSmoothVectorBundleCore {𝕜 : Type _} [NontriviallyNormedField 𝕜] {E : Type _} - [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type _} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) (M : Type _) [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] (F : Type _) [NormedAddCommGroup F] [NormedSpace 𝕜 F] : - BasicSmoothVectorBundleCore I M F - where - coordChange i j x := ContinuousLinearMap.id 𝕜 F - coordChange_self i x hx v := rfl - coordChange_comp i j k x hx v := rfl - coordChange_smooth_clm i j := by - dsimp - exact contDiffOn_const -#align trivial_basic_smooth_vector_bundle_core trivialBasicSmoothVectorBundleCore - -namespace BasicSmoothVectorBundleCore - -variable {𝕜 : Type _} [NontriviallyNormedField 𝕜] {E : Type _} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {H : Type _} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type _} - [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] {F : Type _} - [NormedAddCommGroup F] [NormedSpace 𝕜 F] (Z : BasicSmoothVectorBundleCore I M F) - -instance : Inhabited (BasicSmoothVectorBundleCore I M F) := - ⟨trivialBasicSmoothVectorBundleCore I M F⟩ - -/-- A reformulation of `coord_change_comp`, formulated in terms of a point in `M`. -The conditions on this point a significantly simpler than in `coord_change_comp`. -/ -theorem coordChange_comp' {i j k : atlas H M} {x : M} (hi : x ∈ i.1.source) (hj : x ∈ j.1.source) - (hk : x ∈ k.1.source) (v : F) : - Z.coordChange j k (j x) (Z.coordChange i j (i x) v) = Z.coordChange i k (i x) v := - by - rw [show j x = _ by rw [← i.1.left_inv hi]] - apply Z.coord_change_comp - simp only [hi, hj, hk, mfld_simps] -#align basic_smooth_vector_bundle_core.coord_change_comp' BasicSmoothVectorBundleCore.coordChange_comp' - -/-- A reformulation of `coord_change_self`, formulated in terms of a point in `M`. -/ -theorem coordChange_self' {i : atlas H M} {x : M} (hi : x ∈ i.1.source) (v : F) : - Z.coordChange i i (i x) v = v := - Z.coordChange_self i (i x) (i.1.MapsTo hi) v -#align basic_smooth_vector_bundle_core.coord_change_self' BasicSmoothVectorBundleCore.coordChange_self' - -/-- `Z.coord_change j i` is a partial inverse of `Z.coord_change i j`. -/ -theorem coordChange_comp_eq_self (i j : atlas H M) {x : H} (hx : x ∈ (i.1.symm.trans j.1).source) - (v : F) : Z.coordChange j i (i.1.symm.trans j.1 x) (Z.coordChange i j x v) = v := - by - rw [Z.coord_change_comp i j i x _ v, Z.coord_change_self _ _ hx.1] - simp only [mfld_simps] at hx - simp only [hx.1, hx.2, mfld_simps] -#align basic_smooth_vector_bundle_core.coord_change_comp_eq_self BasicSmoothVectorBundleCore.coordChange_comp_eq_self - -/-- `Z.coord_change j i` is a partial inverse of `Z.coord_change i j`, -formulated in terms of a point in `M`. -/ -theorem coordChange_comp_eq_self' {i j : atlas H M} {x : M} (hi : x ∈ i.1.source) - (hj : x ∈ j.1.source) (v : F) : Z.coordChange j i (j x) (Z.coordChange i j (i x) v) = v := by - rw [Z.coord_change_comp' hi hj hi v, Z.coord_change_self' hi] -#align basic_smooth_vector_bundle_core.coord_change_comp_eq_self' BasicSmoothVectorBundleCore.coordChange_comp_eq_self' - -theorem coordChange_continuous (i j : atlas H M) : - ContinuousOn (Z.coordChange i j) (i.1.symm.trans j.1).source := - by - intro x hx - apply - (((Z.coord_change_smooth_clm i j).ContinuousOn.ContinuousWithinAt (mem_image_of_mem I hx)).comp - I.continuous_within_at _).congr - · intro y hy - simp only [mfld_simps] - · simp only [mfld_simps] - · exact maps_to_image I _ -#align basic_smooth_vector_bundle_core.coord_change_continuous BasicSmoothVectorBundleCore.coordChange_continuous - -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -theorem coordChange_smooth (i j : atlas H M) : - ContDiffOn 𝕜 ∞ (fun p : E × F => Z.coordChange i j (I.symm p.1) p.2) - ((I '' (i.1.symm.trans j.1).source) ×ˢ univ) := - by - have A : ContDiff 𝕜 ∞ fun p : (F →L[𝕜] F) × F => p.1 p.2 := - by - apply IsBoundedBilinearMap.contDiff - exact isBoundedBilinearMapApply - have B : - ContDiffOn 𝕜 ∞ (fun p : E × F => (Z.coord_change i j (I.symm p.1), p.snd)) - ((I '' (i.1.symm.trans j.1).source) ×ˢ univ) := - by - apply ContDiffOn.prod _ _ - · - exact - (Z.coord_change_smooth_clm i j).comp cont_diff_fst.cont_diff_on - (prod_subset_preimage_fst _ _) - · exact is_bounded_linear_map.snd.cont_diff.cont_diff_on - exact A.comp_cont_diff_on B -#align basic_smooth_vector_bundle_core.coord_change_smooth BasicSmoothVectorBundleCore.coordChange_smooth - -/-- Vector bundle core associated to a basic smooth bundle core -/ -@[simps coordChange indexAt] -def toVectorBundleCore : VectorBundleCore 𝕜 M F (atlas H M) - where - baseSet i := i.1.source - isOpen_baseSet i := i.1.open_source - indexAt := achart H - mem_baseSet_at x := mem_chart_source H x - coordChange i j x := Z.coordChange i j (i.1 x) - coordChange_self i x hx v := Z.coordChange_self i (i.1 x) (i.1.map_source hx) v - coordChange_comp := fun i j k x ⟨⟨hx1, hx2⟩, hx3⟩ v => - by - have := Z.coord_change_comp i j k (i.1 x) _ v - convert this using 2 - · simp only [hx1, mfld_simps] - · simp only [hx1, hx2, hx3, mfld_simps] - continuousOn_coordChange i j := - by - refine' ((Z.coord_change_continuous i j).comp' i.1.ContinuousOn).mono _ - rintro p ⟨hp₁, hp₂⟩ - refine' ⟨hp₁, i.1.MapsTo hp₁, _⟩ - simp only [i.1.left_inv hp₁, hp₂, mfld_simps] -#align basic_smooth_vector_bundle_core.to_vector_bundle_core BasicSmoothVectorBundleCore.toVectorBundleCore - -@[simp, mfld_simps] -theorem baseSet (i : atlas H M) : (Z.toVectorBundleCore.localTriv i).baseSet = i.1.source := - rfl -#align basic_smooth_vector_bundle_core.base_set BasicSmoothVectorBundleCore.baseSet - -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -@[simp, mfld_simps] -theorem target (i : atlas H M) : (Z.toVectorBundleCore.localTriv i).target = i.1.source ×ˢ univ := - rfl -#align basic_smooth_vector_bundle_core.target BasicSmoothVectorBundleCore.target - -/-- Local chart for the total space of a basic smooth bundle -/ -def chart {e : LocalHomeomorph M H} (he : e ∈ atlas H M) : - LocalHomeomorph Z.toVectorBundleCore.TotalSpace (ModelProd H F) := - (Z.toVectorBundleCore.localTriv ⟨e, he⟩).toLocalHomeomorph.trans - (LocalHomeomorph.prod e (LocalHomeomorph.refl F)) -#align basic_smooth_vector_bundle_core.chart BasicSmoothVectorBundleCore.chart - -theorem chart_apply {x : M} (z : Z.toVectorBundleCore.TotalSpace) : - Z.chart (chart_mem_atlas H x) z = - (chartAt H x z.proj, - Z.coordChange (achart H z.proj) (achart H x) (achart H z.proj z.proj) z.2) := - rfl -#align basic_smooth_vector_bundle_core.chart_apply BasicSmoothVectorBundleCore.chart_apply - -@[simp, mfld_simps] -theorem chart_source (e : LocalHomeomorph M H) (he : e ∈ atlas H M) : - (Z.chart he).source = Z.toVectorBundleCore.proj ⁻¹' e.source := - by - simp only [chart, mem_prod] - mfld_set_tac -#align basic_smooth_vector_bundle_core.chart_source BasicSmoothVectorBundleCore.chart_source - -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -@[simp, mfld_simps] -theorem chart_target (e : LocalHomeomorph M H) (he : e ∈ atlas H M) : - (Z.chart he).target = e.target ×ˢ univ := - by - simp only [chart] - mfld_set_tac -#align basic_smooth_vector_bundle_core.chart_target BasicSmoothVectorBundleCore.chart_target - -/-- The total space of a basic smooth bundle is endowed with a charted space structure, where the -charts are in bijection with the charts of the basis. -/ -@[simps (config := lemmasOnly) chartAt] -instance toChartedSpace : ChartedSpace (ModelProd H F) Z.toVectorBundleCore.TotalSpace - where - atlas := ⋃ (e : LocalHomeomorph M H) (he : e ∈ atlas H M), {Z.chart he} - chartAt p := Z.chart (chart_mem_atlas H p.1) - mem_chart_source p := by simp [mem_chart_source] - chart_mem_atlas p := - by - simp only [mem_Union, mem_singleton_iff, chart_mem_atlas] - exact ⟨chart_at H p.1, chart_mem_atlas H p.1, rfl⟩ -#align basic_smooth_vector_bundle_core.to_charted_space BasicSmoothVectorBundleCore.toChartedSpace - -theorem mem_atlas_iff (f : LocalHomeomorph Z.toVectorBundleCore.TotalSpace (ModelProd H F)) : - f ∈ atlas (ModelProd H F) Z.toVectorBundleCore.TotalSpace ↔ - ∃ (e : LocalHomeomorph M H)(he : e ∈ atlas H M), f = Z.chart he := - by simp only [atlas, mem_Union, mem_singleton_iff] -#align basic_smooth_vector_bundle_core.mem_atlas_iff BasicSmoothVectorBundleCore.mem_atlas_iff - -@[simp, mfld_simps] -theorem mem_chart_source_iff (p q : Z.toVectorBundleCore.TotalSpace) : - p ∈ (chartAt (ModelProd H F) q).source ↔ p.1 ∈ (chartAt H q.1).source := by - simp only [chart_at, mfld_simps] -#align basic_smooth_vector_bundle_core.mem_chart_source_iff BasicSmoothVectorBundleCore.mem_chart_source_iff - -@[simp, mfld_simps] -theorem mem_chart_target_iff (p : H × F) (q : Z.toVectorBundleCore.TotalSpace) : - p ∈ (chartAt (ModelProd H F) q).target ↔ p.1 ∈ (chartAt H q.1).target := by - simp only [chart_at, mfld_simps] -#align basic_smooth_vector_bundle_core.mem_chart_target_iff BasicSmoothVectorBundleCore.mem_chart_target_iff - -@[simp, mfld_simps] -theorem coe_chartAt_fst (p q : Z.toVectorBundleCore.TotalSpace) : - ((chartAt (ModelProd H F) q) p).1 = chartAt H q.1 p.1 := - rfl -#align basic_smooth_vector_bundle_core.coe_chart_at_fst BasicSmoothVectorBundleCore.coe_chartAt_fst - -@[simp, mfld_simps] -theorem coe_chartAt_symm_fst (p : H × F) (q : Z.toVectorBundleCore.TotalSpace) : - ((chartAt (ModelProd H F) q).symm p).1 = ((chartAt H q.1).symm : H → M) p.1 := - rfl -#align basic_smooth_vector_bundle_core.coe_chart_at_symm_fst BasicSmoothVectorBundleCore.coe_chartAt_symm_fst - -/- ./././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 -/ -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -/-- Smooth manifold structure on the total space of a basic smooth bundle -/ -instance to_smooth_manifold : - SmoothManifoldWithCorners (I.Prod 𝓘(𝕜, F)) Z.toVectorBundleCore.TotalSpace := - by - /- We have to check that the charts belong to the smooth groupoid, i.e., they are smooth on their - source, and their inverses are smooth on the target. Since both objects are of the same kind, it - suffices to prove the first statement in A below, and then glue back the pieces at the end. -/ - let J := ModelWithCorners.toLocalEquiv (I.prod 𝓘(𝕜, F)) - have A : - ∀ (e e' : LocalHomeomorph M H) (he : e ∈ atlas H M) (he' : e' ∈ atlas H M), - ContDiffOn 𝕜 ∞ (J ∘ (Z.chart he).symm.trans (Z.chart he') ∘ J.symm) - (J.symm ⁻¹' ((Z.chart he).symm.trans (Z.chart he')).source ∩ range J) := - by - intro e e' he he' - have : - J.symm ⁻¹' ((chart Z he).symm.trans (chart Z he')).source ∩ range J = - (I.symm ⁻¹' (e.symm.trans e').source ∩ range I) ×ˢ univ := - by - simp only [J, chart, ModelWithCorners.prod] - mfld_set_tac - rw [this] - -- check separately that the two components of the coordinate change are smooth - apply ContDiffOn.prod - show - ContDiffOn 𝕜 ∞ (fun p : E × F => (I ∘ e' ∘ e.symm ∘ I.symm) p.1) - ((I.symm ⁻¹' (e.symm.trans e').source ∩ range I) ×ˢ univ) - · -- the coordinate change on the base is just a coordinate change for `M`, smooth since - -- `M` is smooth - have A : - ContDiffOn 𝕜 ∞ (I ∘ e.symm.trans e' ∘ I.symm) - (I.symm ⁻¹' (e.symm.trans e').source ∩ range I) := - (HasGroupoid.compatible (contDiffGroupoid ∞ I) he he').1 - have B : - ContDiffOn 𝕜 ∞ (fun p : E × F => p.1) - ((I.symm ⁻¹' (e.symm.trans e').source ∩ range I) ×ˢ univ) := - cont_diff_fst.cont_diff_on - exact ContDiffOn.comp A B (prod_subset_preimage_fst _ _) - show - ContDiffOn 𝕜 ∞ - (fun p : E × F => - Z.coord_change ⟨chart_at H (e.symm (I.symm p.1)), _⟩ ⟨e', he'⟩ - ((chart_at H (e.symm (I.symm p.1)) : M → H) (e.symm (I.symm p.1))) - (Z.coord_change ⟨e, he⟩ ⟨chart_at H (e.symm (I.symm p.1)), _⟩ (e (e.symm (I.symm p.1))) - p.2)) - ((I.symm ⁻¹' (e.symm.trans e').source ∩ range I) ×ˢ univ) - · /- The coordinate change in the fiber is more complicated as its definition involves the - reference chart chosen at each point. However, it appears with its inverse, so using the - cocycle property one can get rid of it, and then conclude using the smoothness of the - cocycle as given in the definition of basic smooth bundles. -/ - have := Z.coord_change_smooth ⟨e, he⟩ ⟨e', he'⟩ - rw [I.image_eq] at this - apply ContDiffOn.congr this - rintro ⟨x, v⟩ hx - simp only [mfld_simps] at hx - let f := chart_at H (e.symm (I.symm x)) - have A : I.symm x ∈ ((e.symm.trans f).trans (f.symm.trans e')).source := by - simp only [hx.1.1, hx.1.2, mfld_simps] - rw [e.right_inv hx.1.1] - have := Z.coord_change_comp ⟨e, he⟩ ⟨f, chart_mem_atlas _ _⟩ ⟨e', he'⟩ (I.symm x) A v - simpa only using this - refine' @SmoothManifoldWithCorners.mk _ _ _ _ _ _ _ _ _ _ _ ⟨_⟩ - intro e₀ e₀' he₀ he₀' - rcases(Z.mem_atlas_iff _).1 he₀ with ⟨e, he, rfl⟩ - rcases(Z.mem_atlas_iff _).1 he₀' with ⟨e', he', rfl⟩ - rw [contDiffGroupoid, mem_groupoid_of_pregroupoid] - exact ⟨A e e' he he', A e' e he' he⟩ -#align basic_smooth_vector_bundle_core.to_smooth_manifold BasicSmoothVectorBundleCore.to_smooth_manifold - -end BasicSmoothVectorBundleCore - -section TangentBundle - -variable {𝕜 : Type _} [NontriviallyNormedField 𝕜] {E : Type _} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {H : Type _} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) (M : Type _) - [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] - -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ -/-- Basic smooth bundle core version of the tangent bundle of a smooth manifold `M` modelled over a -model with corners `I` on `(E, H)`. The fibers are equal to `E`, and the coordinate change in the -fiber corresponds to the derivative of the coordinate change in `M`. -/ -@[simps] -def tangentBundleCore : BasicSmoothVectorBundleCore I M E - where - coordChange i j x := fderivWithin 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) (range I) (I x) - coordChange_smooth_clm i j := by - rw [I.image_eq] - have A : - ContDiffOn 𝕜 ∞ (I ∘ i.1.symm.trans j.1 ∘ I.symm) - (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) := - (HasGroupoid.compatible (contDiffGroupoid ∞ I) i.2 j.2).1 - have B : UniqueDiffOn 𝕜 (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) := - I.unique_diff_preimage_source - have C : - ContDiffOn 𝕜 ∞ - (fun p : E × E => - (fderivWithin 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) p.1 : - E → E) - p.2) - ((I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) ×ˢ univ) := - contDiffOn_fderivWithin_apply A B le_top - have D : - ∀ x ∈ I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I, - fderivWithin 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) (range I) x = - fderivWithin 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) x := - by - intro x hx - have N : I.symm ⁻¹' (i.1.symm.trans j.1).source ∈ nhds x := - I.continuous_symm.continuous_at.preimage_mem_nhds - (IsOpen.mem_nhds (LocalHomeomorph.open_source _) hx.1) - symm - rw [inter_comm] - exact fderivWithin_inter N (I.unique_diff _ hx.2) - apply (A.fderiv_within B le_top).congr - intro x hx - simp only [mfld_simps] at hx - simp only [hx, D, mfld_simps] - coordChange_self i x hx v := - by - /- Locally, a self-change of coordinate is just the identity, thus its derivative is the - identity. One just needs to write this carefully, paying attention to the sets where the - functions are defined. -/ - have A : I.symm ⁻¹' (i.1.symm.trans i.1).source ∩ range I ∈ 𝓝[range I] I x := - by - rw [inter_comm] - apply inter_mem_nhdsWithin - apply - I.continuous_symm.continuous_at.preimage_mem_nhds - (IsOpen.mem_nhds (LocalHomeomorph.open_source _) _) - simp only [hx, i.1.map_target, mfld_simps] - have B : ∀ᶠ y in 𝓝[range I] I x, (I ∘ i.1 ∘ i.1.symm ∘ I.symm) y = (id : E → E) y := - by - filter_upwards [A]with _ hy - rw [← I.image_eq] at hy - rcases hy with ⟨z, hz⟩ - simp only [mfld_simps] at hz - simp only [hz.2.symm, hz.1, mfld_simps] - have C : - fderivWithin 𝕜 (I ∘ i.1 ∘ i.1.symm ∘ I.symm) (range I) (I x) = - fderivWithin 𝕜 (id : E → E) (range I) (I x) := - Filter.EventuallyEq.fderivWithin_eq I.unique_diff_at_image B (by simp only [hx, mfld_simps]) - rw [fderivWithin_id I.unique_diff_at_image] at C - rw [C] - rfl - coordChange_comp i j u x hx := - by - /- The cocycle property is just the fact that the derivative of a composition is the product of - the derivatives. One needs however to check that all the functions one considers are smooth, and - to pay attention to the domains where these functions are defined, making this proof a little - bit cumbersome although there is nothing complicated here. -/ - have M : I x ∈ I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I := - ⟨by simpa only [mem_preimage, ModelWithCorners.left_inv] using hx, mem_range_self _⟩ - have U : - UniqueDiffWithinAt 𝕜 - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) (I x) := - I.unique_diff_preimage_source _ M - have A : - fderivWithin 𝕜 ((I ∘ u.1 ∘ j.1.symm ∘ I.symm) ∘ I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) (I x) = - (fderivWithin 𝕜 (I ∘ u.1 ∘ j.1.symm ∘ I.symm) - (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I) - ((I ∘ j.1 ∘ i.1.symm ∘ I.symm) (I x))).comp - (fderivWithin 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) - (I x)) := - by - apply fderivWithin.comp _ _ _ _ U - show - DifferentiableWithinAt 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) (I x) - · have A : - ContDiffOn 𝕜 ∞ (I ∘ i.1.symm.trans j.1 ∘ I.symm) - (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) := - (HasGroupoid.compatible (contDiffGroupoid ∞ I) i.2 j.2).1 - have B : - DifferentiableOn 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) := - by - apply (A.differentiable_on le_top).mono - have : - ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ⊆ - (i.1.symm.trans j.1).source := - inter_subset_left _ _ - exact inter_subset_inter (preimage_mono this) (subset.refl (range I)) - apply B - simpa only [mfld_simps] using hx - show - DifferentiableWithinAt 𝕜 (I ∘ u.1 ∘ j.1.symm ∘ I.symm) - (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I) ((I ∘ j.1 ∘ i.1.symm ∘ I.symm) (I x)) - · have A : - ContDiffOn 𝕜 ∞ (I ∘ j.1.symm.trans u.1 ∘ I.symm) - (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I) := - (HasGroupoid.compatible (contDiffGroupoid ∞ I) j.2 u.2).1 - apply A.differentiable_on le_top - rw [LocalHomeomorph.trans_source] at hx - simp only [mfld_simps] - exact hx.2 - show - I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I ⊆ - I ∘ j.1 ∘ i.1.symm ∘ I.symm ⁻¹' (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I) - · intro y hy - simp only [mfld_simps] at hy - rw [LocalHomeomorph.left_inv] at hy - · simp only [hy, mfld_simps] - · exact hy.1.1.2 - have B : - fderivWithin 𝕜 ((I ∘ u.1 ∘ j.1.symm ∘ I.symm) ∘ I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) (I x) = - fderivWithin 𝕜 (I ∘ u.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) (I x) := - haveI E : - ∀ y ∈ I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I, - ((I ∘ u.1 ∘ j.1.symm ∘ I.symm) ∘ I ∘ j.1 ∘ i.1.symm ∘ I.symm) y = - (I ∘ u.1 ∘ i.1.symm ∘ I.symm) y := - by - intro y hy - simp only [Function.comp_apply, ModelWithCorners.left_inv] - rw [j.1.left_inv] - exact hy.1.1.2 - fderivWithin_congr U E (E _ M) - have C : - fderivWithin 𝕜 (I ∘ u.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) (I x) = - fderivWithin 𝕜 (I ∘ u.1 ∘ i.1.symm ∘ I.symm) (range I) (I x) := - by - rw [inter_comm] - apply fderivWithin_inter _ I.unique_diff_at_image - apply - I.continuous_symm.continuous_at.preimage_mem_nhds - (IsOpen.mem_nhds (LocalHomeomorph.open_source _) _) - simpa only [ModelWithCorners.left_inv] using hx - have D : - fderivWithin 𝕜 (I ∘ u.1 ∘ j.1.symm ∘ I.symm) - (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I) ((I ∘ j.1 ∘ i.1.symm ∘ I.symm) (I x)) = - fderivWithin 𝕜 (I ∘ u.1 ∘ j.1.symm ∘ I.symm) (range I) - ((I ∘ j.1 ∘ i.1.symm ∘ I.symm) (I x)) := - by - rw [inter_comm] - apply fderivWithin_inter _ I.unique_diff_at_image - apply - I.continuous_symm.continuous_at.preimage_mem_nhds - (IsOpen.mem_nhds (LocalHomeomorph.open_source _) _) - rw [LocalHomeomorph.trans_source] at hx - simp only [mfld_simps] - exact hx.2 - have E : - fderivWithin 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) (I x) = - fderivWithin 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) (range I) (I x) := - by - rw [inter_comm] - apply fderivWithin_inter _ I.unique_diff_at_image - apply - I.continuous_symm.continuous_at.preimage_mem_nhds - (IsOpen.mem_nhds (LocalHomeomorph.open_source _) _) - simpa only [ModelWithCorners.left_inv] using hx - rw [B, C, D, E] at A - simp only [A, ContinuousLinearMap.coe_comp', mfld_simps] -#align tangent_bundle_core tangentBundleCore - -variable {M} - -include I - -/-- The tangent space at a point of the manifold `M`. It is just `E`. We could use instead -`(tangent_bundle_core I M).to_vector_bundle_core.fiber x`, but we use `E` to help the -kernel. --/ -@[nolint unused_arguments] -def TangentSpace (x : M) : Type _ := - E -#align tangent_space TangentSpace - -omit I - -variable (M) - --- is empty if the base manifold is empty -/-- The tangent bundle to a smooth manifold, as a Sigma type. Defined in terms of -`bundle.total_space` to be able to put a suitable topology on it. -/ -@[nolint has_nonempty_instance, reducible] -def TangentBundle := - Bundle.TotalSpace (TangentSpace I : M → Type _) -#align tangent_bundle TangentBundle - --- mathport name: exprTM -local notation "TM" => TangentBundle I M - -/-- The projection from the tangent bundle of a smooth manifold to the manifold. As the tangent -bundle is represented internally as a sigma type, the notation `p.1` also works for the projection -of the point `p`. -/ -def TangentBundle.proj : TM → M := fun p => p.1 -#align tangent_bundle.proj TangentBundle.proj - -variable {M} - -@[simp, mfld_simps] -theorem TangentBundle.proj_apply (x : M) (v : TangentSpace I x) : - TangentBundle.proj I M ⟨x, v⟩ = x := - rfl -#align tangent_bundle.proj_apply TangentBundle.proj_apply - -section TangentBundleInstances - -/- In general, the definition of tangent_bundle and tangent_space are not reducible, so that type -class inference does not pick wrong instances. In this section, we record the right instances for -them, noting in particular that the tangent bundle is a smooth manifold. -/ -section - -attribute [local reducible] TangentSpace - -variable {M} (x : M) - -instance : TopologicalSpace (TangentSpace I x) := by infer_instance - -instance : AddCommGroup (TangentSpace I x) := by infer_instance - -instance : TopologicalAddGroup (TangentSpace I x) := by infer_instance - -instance : Module 𝕜 (TangentSpace I x) := by infer_instance - -instance : Inhabited (TangentSpace I x) := - ⟨0⟩ - -end - -variable (M) - -instance : TopologicalSpace TM := - (tangentBundleCore I M).toVectorBundleCore.toTopologicalSpace - -instance : ChartedSpace (ModelProd H E) TM := - (tangentBundleCore I M).toChartedSpace - -instance : SmoothManifoldWithCorners I.tangent TM := - (tangentBundleCore I M).to_smooth_manifold - -instance : FiberBundle E (TangentSpace I : M → Type _) := - (tangentBundleCore I M).toVectorBundleCore.FiberBundle - -instance : VectorBundle 𝕜 E (TangentSpace I : M → Type _) := - (tangentBundleCore I M).toVectorBundleCore.VectorBundle - -end TangentBundleInstances - -variable (M) - -/-- The tangent bundle projection on the basis is a continuous map. -/ -theorem tangentBundle_proj_continuous : Continuous (TangentBundle.proj I M) := - (tangentBundleCore I M).toVectorBundleCore.continuous_proj -#align tangent_bundle_proj_continuous tangentBundle_proj_continuous - -/-- The tangent bundle projection on the basis is an open map. -/ -theorem tangentBundle_proj_open : IsOpenMap (TangentBundle.proj I M) := - (tangentBundleCore I M).toVectorBundleCore.isOpenMap_proj -#align tangent_bundle_proj_open tangentBundle_proj_open - -/-- In the tangent bundle to the model space, the charts are just the canonical identification -between a product type and a sigma type, a.k.a. `equiv.sigma_equiv_prod`. -/ -@[simp, mfld_simps] -theorem tangentBundle_model_space_chartAt (p : TangentBundle I H) : - (chartAt (ModelProd H E) p).toLocalEquiv = (Equiv.sigmaEquivProd H E).toLocalEquiv := - by - have A : ∀ x_fst, fderivWithin 𝕜 (I ∘ I.symm) (range I) (I x_fst) = ContinuousLinearMap.id 𝕜 E := - by - intro x_fst - have : - fderivWithin 𝕜 (I ∘ I.symm) (range I) (I x_fst) = fderivWithin 𝕜 id (range I) (I x_fst) := - by - refine' fderivWithin_congr I.unique_diff_at_image (fun y hy => _) (by simp) - exact ModelWithCorners.right_inv _ hy - rwa [fderivWithin_id I.unique_diff_at_image] at this - ext x : 1 - show - (chart_at (ModelProd H E) p : TangentBundle I H → ModelProd H E) x = - (Equiv.sigmaEquivProd H E) x - · cases x - simp only [chart_at, BasicSmoothVectorBundleCore.chart, tangentBundleCore, - BasicSmoothVectorBundleCore.toVectorBundleCore, A, Prod.mk.inj_iff, - ContinuousLinearMap.coe_id', mfld_simps] - show ∀ x, (chart_at (ModelProd H E) p).toLocalEquiv.symm x = (Equiv.sigmaEquivProd H E).symm x - · rintro ⟨x_fst, x_snd⟩ - simp only [BasicSmoothVectorBundleCore.toVectorBundleCore, tangentBundleCore, A, - ContinuousLinearMap.coe_id', BasicSmoothVectorBundleCore.chart, chart_at, - ContinuousLinearMap.coe_coe, Sigma.mk.inj_iff, mfld_simps] - show (chart_at (ModelProd H E) p).toLocalEquiv.source = univ - · simp only [chart_at, mfld_simps] -#align tangent_bundle_model_space_chart_at tangentBundle_model_space_chartAt - -@[simp, mfld_simps] -theorem tangentBundle_model_space_coe_chartAt (p : TangentBundle I H) : - ⇑(chartAt (ModelProd H E) p) = Equiv.sigmaEquivProd H E := - by - unfold_coes - simp only [mfld_simps] -#align tangent_bundle_model_space_coe_chart_at tangentBundle_model_space_coe_chartAt - -@[simp, mfld_simps] -theorem tangentBundle_model_space_coe_chartAt_symm (p : TangentBundle I H) : - ((chartAt (ModelProd H E) p).symm : ModelProd H E → TangentBundle I H) = - (Equiv.sigmaEquivProd H E).symm := - by - unfold_coes - simp only [mfld_simps] -#align tangent_bundle_model_space_coe_chart_at_symm tangentBundle_model_space_coe_chartAt_symm - -variable (H) - -/-- The canonical identification between the tangent bundle to the model space and the product, -as a homeomorphism -/ -def tangentBundleModelSpaceHomeomorph : TangentBundle I H ≃ₜ ModelProd H E := - { - Equiv.sigmaEquivProd H - E with - continuous_toFun := by - let p : TangentBundle I H := ⟨I.symm (0 : E), (0 : E)⟩ - have : Continuous (chart_at (ModelProd H E) p) := - by - rw [continuous_iff_continuousOn_univ] - convert LocalHomeomorph.continuousOn _ - simp only [mfld_simps] - simpa only [mfld_simps] using this - continuous_invFun := - by - let p : TangentBundle I H := ⟨I.symm (0 : E), (0 : E)⟩ - have : Continuous (chart_at (ModelProd H E) p).symm := - by - rw [continuous_iff_continuousOn_univ] - convert LocalHomeomorph.continuousOn _ - simp only [mfld_simps] - simpa only [mfld_simps] using this } -#align tangent_bundle_model_space_homeomorph tangentBundleModelSpaceHomeomorph - -@[simp, mfld_simps] -theorem tangentBundleModelSpaceHomeomorph_coe : - (tangentBundleModelSpaceHomeomorph H I : TangentBundle I H → ModelProd H E) = - Equiv.sigmaEquivProd H E := - rfl -#align tangent_bundle_model_space_homeomorph_coe tangentBundleModelSpaceHomeomorph_coe - -@[simp, mfld_simps] -theorem tangentBundleModelSpaceHomeomorph_coe_symm : - ((tangentBundleModelSpaceHomeomorph H I).symm : ModelProd H E → TangentBundle I H) = - (Equiv.sigmaEquivProd H E).symm := - rfl -#align tangent_bundle_model_space_homeomorph_coe_symm tangentBundleModelSpaceHomeomorph_coe_symm - -end TangentBundle - diff --git a/Mathbin/Geometry/Manifold/VectorBundle/Basic.lean b/Mathbin/Geometry/Manifold/VectorBundle/Basic.lean index 8a672120c5..32cc4ea7e0 100644 --- a/Mathbin/Geometry/Manifold/VectorBundle/Basic.lean +++ b/Mathbin/Geometry/Manifold/VectorBundle/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Heather Macbeth ! This file was ported from Lean 3 source module geometry.manifold.vector_bundle.basic -! leanprover-community/mathlib commit ddec54a71a0dd025c05445d467f1a2b7d586a3ba +! leanprover-community/mathlib commit 0187644979f2d3e10a06e916a869c994facd9a87 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -62,7 +62,11 @@ assert_not_exists mfderiv open Bundle Set LocalHomeomorph -open Manifold Bundle +open Function (id_def) + +open Filter + +open Manifold Bundle Topology variable {𝕜 B B' F M : Type _} {E : B → Type _} @@ -107,6 +111,148 @@ theorem FiberBundle.chartedSpace_chartAt (x : TotalSpace E) : rw [Trivialization.coe_coe, Trivialization.coe_fst' _ (mem_base_set_trivialization_at F E x.proj)] #align fiber_bundle.charted_space_chart_at FiberBundle.chartedSpace_chartAt +theorem FiberBundle.chartedSpace_chartAt_symm_fst (x : TotalSpace E) (y : ModelProd HB F) + (hy : y ∈ (chartAt (ModelProd HB F) x).target) : + ((chartAt (ModelProd HB F) x).symm y).proj = (chartAt HB x.proj).symm y.1 := + by + simp only [FiberBundle.chartedSpace_chartAt, mfld_simps] at hy⊢ + exact (trivialization_at F E x.proj).proj_symm_apply hy.2 +#align fiber_bundle.charted_space_chart_at_symm_fst FiberBundle.chartedSpace_chartAt_symm_fst + +end + +section + +variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] + [TopologicalSpace (TotalSpace E)] [∀ x, TopologicalSpace (E x)] {EB : Type _} + [NormedAddCommGroup EB] [NormedSpace 𝕜 EB] {HB : Type _} [TopologicalSpace HB] + (IB : ModelWithCorners 𝕜 EB HB) (E' : B → Type _) [∀ x, Zero (E' x)] {EM : Type _} + [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type _} [TopologicalSpace HM] + {IM : ModelWithCorners 𝕜 EM HM} [TopologicalSpace M] [ChartedSpace HM M] + [Is : SmoothManifoldWithCorners IM M] {n : ℕ∞} + +variable [TopologicalSpace B] [ChartedSpace HB B] [FiberBundle F E] + +protected theorem FiberBundle.extChartAt (x : TotalSpace E) : + extChartAt (IB.Prod 𝓘(𝕜, F)) x = + (trivializationAt F E x.proj).toLocalEquiv ≫ + (extChartAt IB x.proj).Prod (LocalEquiv.refl F) := + by + simp_rw [extChartAt, FiberBundle.chartedSpace_chartAt, extend] + simp only [LocalEquiv.trans_assoc, mfld_simps] +#align fiber_bundle.ext_chart_at FiberBundle.extChartAt + +/-! ### Smoothness of maps in/out fiber bundles + +Note: For these results we don't need that the bundle is a smooth vector bundle, or even a vector +bundle at all, just that it is a fiber bundle over a charted base space. +-/ + + +namespace Bundle + +variable {F E IB} + +/-- Characterization of C^n functions into a smooth vector bundle. -/ +theorem contMdiffWithinAt_totalSpace (f : M → TotalSpace E) {s : Set M} {x₀ : M} : + ContMdiffWithinAt IM (IB.Prod 𝓘(𝕜, F)) n f s x₀ ↔ + ContMdiffWithinAt IM IB n (fun x => (f x).proj) s x₀ ∧ + ContMdiffWithinAt IM 𝓘(𝕜, F) n (fun x => (trivializationAt F E (f x₀).proj (f x)).2) s x₀ := + by + simp (config := { singlePass := true }) only [contMdiffWithinAt_iff_target] + rw [and_and_and_comm, ← continuous_within_at_total_space, and_congr_right_iff] + intro hf + simp_rw [modelWithCornersSelf_prod, FiberBundle.extChartAt, Function.comp, LocalEquiv.trans_apply, + LocalEquiv.prod_coe, LocalEquiv.refl_coe, extChartAt_self_apply, modelWithCornersSelf_coe, + id_def] + refine' (contMdiffWithinAt_prod_iff _).trans _ + -- rw doesn't do this? + have h1 : (fun x => (f x).proj) ⁻¹' (trivialization_at F E (f x₀).proj).baseSet ∈ 𝓝[s] x₀ := + ((continuous_proj F E).ContinuousWithinAt.comp hf (maps_to_image f s)).preimage_mem_nhdsWithin + ((Trivialization.open_baseSet _).mem_nhds (mem_base_set_trivialization_at F E _)) + refine' + and_congr (eventually_eq.cont_mdiff_within_at_iff (eventually_of_mem h1 fun x hx => _) _) + Iff.rfl + · simp_rw [Function.comp, LocalHomeomorph.coe_coe, Trivialization.coe_coe] + rw [Trivialization.coe_fst'] + exact hx + · simp only [mfld_simps] +#align bundle.cont_mdiff_within_at_total_space Bundle.contMdiffWithinAt_totalSpace + +/-- Characterization of C^n functions into a smooth vector bundle. -/ +theorem contMdiffAt_totalSpace (f : M → TotalSpace E) (x₀ : M) : + ContMdiffAt IM (IB.Prod 𝓘(𝕜, F)) n f x₀ ↔ + ContMdiffAt IM IB n (fun x => (f x).proj) x₀ ∧ + ContMdiffAt IM 𝓘(𝕜, F) n (fun x => (trivializationAt F E (f x₀).proj (f x)).2) x₀ := + by + simp_rw [← contMdiffWithinAt_univ] + exact cont_mdiff_within_at_total_space f +#align bundle.cont_mdiff_at_total_space Bundle.contMdiffAt_totalSpace + +variable (E) + +theorem contMdiff_proj : ContMdiff (IB.Prod 𝓘(𝕜, F)) IB n (π E) := + by + intro x + rw [ContMdiffAt, contMdiffWithinAt_iff'] + refine' ⟨(continuous_proj F E).ContinuousWithinAt, _⟩ + simp_rw [(· ∘ ·), FiberBundle.extChartAt] + apply cont_diff_within_at_fst.congr + · rintro ⟨a, b⟩ hab + simp only [mfld_simps] at hab + have : ((chart_at HB x.1).symm (IB.symm a), b) ∈ (trivialization_at F E x.fst).target := by + simp only [hab, mfld_simps] + simp only [Trivialization.proj_symm_apply _ this, hab, mfld_simps] + · simp only [mfld_simps] +#align bundle.cont_mdiff_proj Bundle.contMdiff_proj + +theorem smooth_proj : Smooth (IB.Prod 𝓘(𝕜, F)) IB (π E) := + contMdiff_proj E +#align bundle.smooth_proj Bundle.smooth_proj + +theorem contMdiffOn_proj {s : Set (TotalSpace E)} : ContMdiffOn (IB.Prod 𝓘(𝕜, F)) IB n (π E) s := + (Bundle.contMdiff_proj E).ContMdiffOn +#align bundle.cont_mdiff_on_proj Bundle.contMdiffOn_proj + +theorem smoothOn_proj {s : Set (TotalSpace E)} : SmoothOn (IB.Prod 𝓘(𝕜, F)) IB (π E) s := + contMdiffOn_proj E +#align bundle.smooth_on_proj Bundle.smoothOn_proj + +theorem contMdiffAt_proj {p : TotalSpace E} : ContMdiffAt (IB.Prod 𝓘(𝕜, F)) IB n (π E) p := + (Bundle.contMdiff_proj E).ContMdiffAt +#align bundle.cont_mdiff_at_proj Bundle.contMdiffAt_proj + +theorem smoothAt_proj {p : TotalSpace E} : SmoothAt (IB.Prod 𝓘(𝕜, F)) IB (π E) p := + Bundle.contMdiffAt_proj E +#align bundle.smooth_at_proj Bundle.smoothAt_proj + +theorem contMdiffWithinAt_proj {s : Set (TotalSpace E)} {p : TotalSpace E} : + ContMdiffWithinAt (IB.Prod 𝓘(𝕜, F)) IB n (π E) s p := + (Bundle.contMdiffAt_proj E).ContMdiffWithinAt +#align bundle.cont_mdiff_within_at_proj Bundle.contMdiffWithinAt_proj + +theorem smoothWithinAt_proj {s : Set (TotalSpace E)} {p : TotalSpace E} : + SmoothWithinAt (IB.Prod 𝓘(𝕜, F)) IB (π E) s p := + Bundle.contMdiffWithinAt_proj E +#align bundle.smooth_within_at_proj Bundle.smoothWithinAt_proj + +variable (𝕜 E) [∀ x, AddCommMonoid (E x)] [∀ x, Module 𝕜 (E x)] [VectorBundle 𝕜 F E] + +theorem smooth_zeroSection : Smooth IB (IB.Prod 𝓘(𝕜, F)) (zeroSection E) := + by + intro x + rw [Bundle.contMdiffAt_totalSpace] + refine' ⟨contMdiffAt_id, cont_mdiff_at_const.congr_of_eventually_eq _⟩ + · exact 0 + refine' + eventually_of_mem + ((trivialization_at F E x).open_baseSet.mem_nhds (mem_base_set_trivialization_at F E x)) + fun x' hx' => _ + simp_rw [zero_section_proj, (trivialization_at F E x).zeroSection 𝕜 hx'] +#align bundle.smooth_zero_section Bundle.smooth_zeroSection + +end Bundle + end /-! ### Smooth vector bundles -/ @@ -116,7 +262,9 @@ variable [NontriviallyNormedField 𝕜] [∀ x, AddCommMonoid (E x)] [∀ x, Mod [NormedAddCommGroup F] [NormedSpace 𝕜 F] [TopologicalSpace (TotalSpace E)] [∀ x, TopologicalSpace (E x)] {EB : Type _} [NormedAddCommGroup EB] [NormedSpace 𝕜 EB] {HB : Type _} [TopologicalSpace HB] (IB : ModelWithCorners 𝕜 EB HB) [TopologicalSpace B] - [ChartedSpace HB B] [SmoothManifoldWithCorners IB B] + [ChartedSpace HB B] [SmoothManifoldWithCorners IB B] {EM : Type _} [NormedAddCommGroup EM] + [NormedSpace 𝕜 EM] {HM : Type _} [TopologicalSpace HM] {IM : ModelWithCorners 𝕜 EM HM} + [TopologicalSpace M] [ChartedSpace HM M] [Is : SmoothManifoldWithCorners IM M] {n : ℕ∞} variable (F E) [FiberBundle F E] [VectorBundle 𝕜 F E] diff --git a/Mathbin/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathbin/Geometry/Manifold/VectorBundle/Tangent.lean index 38a44a1566..0f39cb69d8 100644 --- a/Mathbin/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathbin/Geometry/Manifold/VectorBundle/Tangent.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Heather Macbeth ! This file was ported from Lean 3 source module geometry.manifold.vector_bundle.tangent -! leanprover-community/mathlib commit ddec54a71a0dd025c05445d467f1a2b7d586a3ba +! leanprover-community/mathlib commit 0187644979f2d3e10a06e916a869c994facd9a87 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -36,7 +36,7 @@ This defines a smooth vector bundle `tangent_bundle` with fibers `tangent_space` open Bundle Set SmoothManifoldWithCorners LocalHomeomorph -open Manifold Topology +open Manifold Topology Bundle noncomputable section @@ -47,10 +47,6 @@ variable {𝕜 : Type _} [NontriviallyNormedField 𝕜] {E : Type _} [NormedAddC variable (I) -namespace hidden - -/- we temporarily put everything in a namespace to not have name clashes with -the existing `tangent_bundle_core`. -/ /-- Auxiliary lemma for tangent spaces: the derivative of a coordinate change between two charts is smooth on its source. -/ theorem contDiffOn_fderiv_coord_change (i j : atlas H M) : @@ -68,7 +64,7 @@ theorem contDiffOn_fderiv_coord_change (i j : atlas H M) : (subset_maximal_atlas I i.2) x hx).mono_of_mem _ exact i.1.extend_coord_change_source_mem_nhdsWithin j.1 I hx -#align hidden.cont_diff_on_fderiv_coord_change hidden.contDiffOn_fderiv_coord_change +#align cont_diff_on_fderiv_coord_change contDiffOn_fderiv_coord_change variable (M) @@ -103,8 +99,8 @@ def tangentBundleCore : VectorBundleCore 𝕜 M E (atlas H M) continuousOn_coordChange i j := by refine' - (cont_diff_on_fderiv_coord_change I i j).ContinuousOn.comp - ((i.1.continuousOn_extend I).mono _) _ + (contDiffOn_fderiv_coord_change I i j).ContinuousOn.comp ((i.1.continuousOn_extend I).mono _) + _ · rw [i.1.extend_source] exact inter_subset_left _ _ simp_rw [← i.1.extend_image_source_inter, maps_to_image] @@ -130,7 +126,7 @@ def tangentBundleCore : VectorBundleCore 𝕜 M E (atlas H M) exact mem_range_self _ · exact I.unique_diff_at_image · rw [Function.comp_apply, i.1.extend_left_inv I hxi] -#align hidden.tangent_bundle_core hidden.tangentBundleCore +#align tangent_bundle_core tangentBundleCore variable {M} @@ -138,7 +134,7 @@ theorem tangentBundleCore_coordChange_achart (x x' z : M) : (tangentBundleCore I M).coordChange (achart H x) (achart H x') z = fderivWithin 𝕜 (extChartAt I x' ∘ (extChartAt I x).symm) (range I) (extChartAt I x z) := rfl -#align hidden.tangent_bundle_core_coord_change_achart hidden.tangentBundleCore_coordChange_achart +#align tangent_bundle_core_coord_change_achart tangentBundleCore_coordChange_achart include I @@ -149,7 +145,7 @@ kernel. @[nolint unused_arguments] def TangentSpace (x : M) : Type _ := E deriving TopologicalSpace, AddCommGroup, TopologicalAddGroup -#align hidden.tangent_space hidden.TangentSpace +#align tangent_space TangentSpace omit I @@ -161,7 +157,7 @@ variable (M) @[nolint has_nonempty_instance, reducible] def TangentBundle := Bundle.TotalSpace (TangentSpace I : M → Type _) -#align hidden.tangent_bundle hidden.TangentBundle +#align tangent_bundle TangentBundle -- mathport name: exprTM local notation "TM" => TangentBundle I M @@ -173,7 +169,7 @@ does not pick wrong instances. In this section, we record the right instances fo them, noting in particular that the tangent bundle is a smooth manifold. -/ section -attribute [local reducible] tangent_space +attribute [local reducible] TangentSpace variable {M} (x : M) @@ -185,43 +181,109 @@ instance : Inhabited (TangentSpace I x) := end instance : TopologicalSpace TM := - (tangentBundleCore I M).toFiberBundleCore.toTopologicalSpace + (tangentBundleCore I M).toTopologicalSpace instance : FiberBundle E (TangentSpace I : M → Type _) := - (tangentBundleCore I M).toFiberBundleCore.FiberBundle + (tangentBundleCore I M).FiberBundle instance : VectorBundle 𝕜 E (TangentSpace I : M → Type _) := (tangentBundleCore I M).VectorBundle -theorem tangentSpace_chartAt (p : TM) : +namespace TangentBundle + +protected theorem chartAt (p : TM) : chartAt (ModelProd H E) p = ((tangentBundleCore I M).toFiberBundleCore.localTriv (achart H p.1)).toLocalHomeomorph ≫ₕ (chartAt H p.1).Prod (LocalHomeomorph.refl E) := rfl -#align hidden.tangent_space_chart_at hidden.tangentSpace_chartAt +#align tangent_bundle.chart_at TangentBundle.chartAt -theorem tangentSpace_chartAt_toLocalEquiv (p : TM) : +theorem chartAt_toLocalEquiv (p : TM) : (chartAt (ModelProd H E) p).toLocalEquiv = (tangentBundleCore I M).toFiberBundleCore.localTrivAsLocalEquiv (achart H p.1) ≫ (chartAt H p.1).toLocalEquiv.Prod (LocalEquiv.refl E) := rfl -#align hidden.tangent_space_chart_at_to_local_equiv hidden.tangentSpace_chartAt_toLocalEquiv +#align tangent_bundle.chart_at_to_local_equiv TangentBundle.chartAt_toLocalEquiv + +theorem trivializationAt_eq_localTriv (x : M) : + trivializationAt E (TangentSpace I) x = + (tangentBundleCore I M).toFiberBundleCore.localTriv (achart H x) := + rfl +#align tangent_bundle.trivialization_at_eq_local_triv TangentBundle.trivializationAt_eq_localTriv + +@[simp, mfld_simps] +theorem trivializationAt_source (x : M) : + (trivializationAt E (TangentSpace I) x).source = π _ ⁻¹' (chartAt H x).source := + rfl +#align tangent_bundle.trivialization_at_source TangentBundle.trivializationAt_source + +/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ +@[simp, mfld_simps] +theorem trivializationAt_target (x : M) : + (trivializationAt E (TangentSpace I) x).target = (chartAt H x).source ×ˢ univ := + rfl +#align tangent_bundle.trivialization_at_target TangentBundle.trivializationAt_target + +@[simp, mfld_simps] +theorem trivializationAt_baseSet (x : M) : + (trivializationAt E (TangentSpace I) x).baseSet = (chartAt H x).source := + rfl +#align tangent_bundle.trivialization_at_base_set TangentBundle.trivializationAt_baseSet + +theorem trivializationAt_apply (x : M) (z : TM) : + trivializationAt E (TangentSpace I) x z = + (z.1, + fderivWithin 𝕜 ((chartAt H x).extend I ∘ ((chartAt H z.1).extend I).symm) (range I) + ((chartAt H z.1).extend I z.1) z.2) := + rfl +#align tangent_bundle.trivialization_at_apply TangentBundle.trivializationAt_apply + +@[simp, mfld_simps] +theorem trivializationAt_fst (x : M) (z : TM) : (trivializationAt E (TangentSpace I) x z).1 = z.1 := + rfl +#align tangent_bundle.trivialization_at_fst TangentBundle.trivializationAt_fst + +@[simp, mfld_simps] +theorem mem_chart_source_iff (p q : TM) : + p ∈ (chartAt (ModelProd H E) q).source ↔ p.1 ∈ (chartAt H q.1).source := by + simp only [FiberBundle.chartedSpace_chartAt, mfld_simps] +#align tangent_bundle.mem_chart_source_iff TangentBundle.mem_chart_source_iff + +@[simp, mfld_simps] +theorem mem_chart_target_iff (p : H × E) (q : TM) : + p ∈ (chartAt (ModelProd H E) q).target ↔ p.1 ∈ (chartAt H q.1).target := by + simp (config := { contextual := true }) only [FiberBundle.chartedSpace_chartAt, + and_iff_left_iff_imp, mfld_simps] +#align tangent_bundle.mem_chart_target_iff TangentBundle.mem_chart_target_iff + +@[simp, mfld_simps] +theorem coe_chartAt_fst (p q : TM) : ((chartAt (ModelProd H E) q) p).1 = chartAt H q.1 p.1 := + rfl +#align tangent_bundle.coe_chart_at_fst TangentBundle.coe_chartAt_fst + +@[simp, mfld_simps] +theorem coe_chartAt_symm_fst (p : H × E) (q : TM) : + ((chartAt (ModelProd H E) q).symm p).1 = ((chartAt H q.1).symm : H → M) p.1 := + rfl +#align tangent_bundle.coe_chart_at_symm_fst TangentBundle.coe_chartAt_symm_fst + +end TangentBundle instance tangentBundleCore.isSmooth : (tangentBundleCore I M).IsSmooth I := by refine' ⟨fun i j => _⟩ rw [SmoothOn, contMdiffOn_iff_source_of_mem_maximalAtlas (subset_maximal_atlas I i.2), contMdiffOn_iff_contDiffOn] - refine' ((cont_diff_on_fderiv_coord_change I i j).congr fun x hx => _).mono _ + refine' ((contDiffOn_fderiv_coord_change I i j).congr fun x hx => _).mono _ · rw [LocalEquiv.trans_source'] at hx - simp_rw [Function.comp_apply, tangent_bundle_core_coord_change, (i.1.extend I).right_inv hx.1] + simp_rw [Function.comp_apply, tangentBundleCore_coordChange, (i.1.extend I).right_inv hx.1] · exact (i.1.extend_image_source_inter j.1 I).Subset · apply inter_subset_left -#align hidden.tangent_bundle_core.is_smooth hidden.tangentBundleCore.isSmooth +#align tangent_bundle_core.is_smooth tangentBundleCore.isSmooth instance TangentBundle.smoothVectorBundle : SmoothVectorBundle E (TangentSpace I : M → Type _) I := (tangentBundleCore I M).SmoothVectorBundle _ -#align hidden.tangent_bundle.smooth_vector_bundle hidden.TangentBundle.smoothVectorBundle +#align tangent_bundle.smooth_vector_bundle TangentBundle.smoothVectorBundle end TangentBundleInstances @@ -237,27 +299,25 @@ theorem tangentBundle_model_space_chartAt (p : TangentBundle I H) : ext x : 1 · ext · rfl - exact - (tangent_bundle_core I H).coordChange_self (achart _ x.1) x.1 (mem_achart_source H x.1) x.2 + exact (tangentBundleCore I H).coordChange_self (achart _ x.1) x.1 (mem_achart_source H x.1) x.2 · intro x ext · rfl apply hEq_of_eq - exact - (tangent_bundle_core I H).coordChange_self (achart _ x.1) x.1 (mem_achart_source H x.1) x.2 - simp_rw [tangent_space_chart_at, FiberBundleCore.localTriv, FiberBundleCore.localTrivAsLocalEquiv, - VectorBundleCore.toFiberBundleCore_baseSet, tangent_bundle_core_base_set] + exact (tangentBundleCore I H).coordChange_self (achart _ x.1) x.1 (mem_achart_source H x.1) x.2 + simp_rw [TangentBundle.chartAt, FiberBundleCore.localTriv, FiberBundleCore.localTrivAsLocalEquiv, + VectorBundleCore.toFiberBundleCore_baseSet, tangentBundleCore_baseSet] simp only [mfld_simps] -#align hidden.tangent_bundle_model_space_chart_at hidden.tangentBundle_model_space_chartAt +#align tangent_bundle_model_space_chart_at tangentBundle_model_space_chartAt @[simp, mfld_simps] theorem tangentBundle_model_space_coe_chartAt (p : TangentBundle I H) : ⇑(chartAt (ModelProd H E) p) = Equiv.sigmaEquivProd H E := by unfold_coes - simp_rw [tangent_bundle_model_space_chart_at] + simp_rw [tangentBundle_model_space_chartAt] rfl -#align hidden.tangent_bundle_model_space_coe_chart_at hidden.tangentBundle_model_space_coe_chartAt +#align tangent_bundle_model_space_coe_chart_at tangentBundle_model_space_coe_chartAt @[simp, mfld_simps] theorem tangentBundle_model_space_coe_chartAt_symm (p : TangentBundle I H) : @@ -265,9 +325,9 @@ theorem tangentBundle_model_space_coe_chartAt_symm (p : TangentBundle I H) : (Equiv.sigmaEquivProd H E).symm := by unfold_coes - simp_rw [LocalHomeomorph.symm_toLocalEquiv, tangent_bundle_model_space_chart_at] + simp_rw [LocalHomeomorph.symm_toLocalEquiv, tangentBundle_model_space_chartAt] rfl -#align hidden.tangent_bundle_model_space_coe_chart_at_symm hidden.tangentBundle_model_space_coe_chartAt_symm +#align tangent_bundle_model_space_coe_chart_at_symm tangentBundle_model_space_coe_chartAt_symm variable (H) @@ -278,37 +338,35 @@ def tangentBundleModelSpaceHomeomorph : TangentBundle I H ≃ₜ ModelProd H E : Equiv.sigmaEquivProd H E with continuous_toFun := by - let p : tangent_bundle I H := ⟨I.symm (0 : E), (0 : E)⟩ + let p : TangentBundle I H := ⟨I.symm (0 : E), (0 : E)⟩ have : Continuous (chart_at (ModelProd H E) p) := by rw [continuous_iff_continuousOn_univ] convert LocalHomeomorph.continuousOn _ - simp only [tangent_space.fiber_bundle, mfld_simps] + simp only [TangentSpace.fiberBundle, mfld_simps] simpa only [mfld_simps] using this continuous_invFun := by - let p : tangent_bundle I H := ⟨I.symm (0 : E), (0 : E)⟩ + let p : TangentBundle I H := ⟨I.symm (0 : E), (0 : E)⟩ have : Continuous (chart_at (ModelProd H E) p).symm := by rw [continuous_iff_continuousOn_univ] convert LocalHomeomorph.continuousOn _ simp only [mfld_simps] simpa only [mfld_simps] using this } -#align hidden.tangent_bundle_model_space_homeomorph hidden.tangentBundleModelSpaceHomeomorph +#align tangent_bundle_model_space_homeomorph tangentBundleModelSpaceHomeomorph @[simp, mfld_simps] theorem tangentBundleModelSpaceHomeomorph_coe : (tangentBundleModelSpaceHomeomorph H I : TangentBundle I H → ModelProd H E) = Equiv.sigmaEquivProd H E := rfl -#align hidden.tangent_bundle_model_space_homeomorph_coe hidden.tangentBundleModelSpaceHomeomorph_coe +#align tangent_bundle_model_space_homeomorph_coe tangentBundleModelSpaceHomeomorph_coe @[simp, mfld_simps] theorem tangentBundleModelSpaceHomeomorph_coe_symm : ((tangentBundleModelSpaceHomeomorph H I).symm : ModelProd H E → TangentBundle I H) = (Equiv.sigmaEquivProd H E).symm := rfl -#align hidden.tangent_bundle_model_space_homeomorph_coe_symm hidden.tangentBundleModelSpaceHomeomorph_coe_symm - -end hidden +#align tangent_bundle_model_space_homeomorph_coe_symm tangentBundleModelSpaceHomeomorph_coe_symm diff --git a/Mathbin/LinearAlgebra/AffineSpace/Basis.lean b/Mathbin/LinearAlgebra/AffineSpace/Basis.lean index f7067faba4..e6ac7b5869 100644 --- a/Mathbin/LinearAlgebra/AffineSpace/Basis.lean +++ b/Mathbin/LinearAlgebra/AffineSpace/Basis.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash ! This file was ported from Lean 3 source module linear_algebra.affine_space.basis -! leanprover-community/mathlib commit 2f4cdce0c2f2f3b8cd58f05d556d03b468e1eb2e +! leanprover-community/mathlib commit 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -196,7 +196,7 @@ theorem coord_apply [DecidableEq ι] (i j : ι) : b.Coord i (b j) = if i = j the @[simp] theorem coord_apply_combination_of_mem (hi : i ∈ s) {w : ι → k} (hw : s.Sum w = 1) : - b.Coord i (s.affineCombination b w) = w i := by + b.Coord i (s.affineCombination k b w) = w i := by classical simp only [coord_apply, hi, Finset.affineCombination_eq_linear_combination, if_true, mul_boole, hw, Function.comp_apply, smul_eq_mul, s.sum_ite_eq, s.map_affine_combination b w hw] @@ -204,7 +204,7 @@ theorem coord_apply_combination_of_mem (hi : i ∈ s) {w : ι → k} (hw : s.Sum @[simp] theorem coord_apply_combination_of_not_mem (hi : i ∉ s) {w : ι → k} (hw : s.Sum w = 1) : - b.Coord i (s.affineCombination b w) = 0 := by + b.Coord i (s.affineCombination k b w) = 0 := by classical simp only [coord_apply, hi, Finset.affineCombination_eq_linear_combination, if_false, mul_boole, hw, Function.comp_apply, smul_eq_mul, s.sum_ite_eq, s.map_affine_combination b w hw] @@ -224,7 +224,7 @@ theorem sum_coord_apply_eq_one [Fintype ι] (q : P) : (∑ i, b.Coord i q) = 1 : @[simp] theorem affineCombination_coord_eq_self [Fintype ι] (q : P) : - (Finset.univ.affineCombination b fun i => b.Coord i q) = q := + (Finset.univ.affineCombination k b fun i => b.Coord i q) = q := by have hq : q ∈ affineSpan k (range b) := by rw [b.tot] @@ -264,7 +264,7 @@ theorem coe_coord_of_subsingleton_eq_one [Subsingleton ι] (i : ι) : (b.Coord i let s : Finset ι := {i} have hi : i ∈ s := by simp have hw : s.sum (Function.const ι (1 : k)) = 1 := by simp - have hq : q = s.affine_combination b (Function.const ι (1 : k)) := by simp + have hq : q = s.affine_combination k b (Function.const ι (1 : k)) := by simp rw [Pi.one_apply, hq, b.coord_apply_combination_of_mem hi hw] #align affine_basis.coe_coord_of_subsingleton_eq_one AffineBasis.coe_coord_of_subsingleton_eq_one @@ -277,7 +277,7 @@ theorem surjective_coord [Nontrivial ι] (i : ι) : Function.Surjective <| b.Coo have hj : j ∈ s := by simp let w : ι → k := fun j' => if j' = i then x else 1 - x have hw : s.sum w = 1 := by simp [hij, Finset.sum_ite, Finset.filter_insert, Finset.filter_eq'] - use s.affine_combination b w + use s.affine_combination k b w simp [b.coord_apply_combination_of_mem hi hw] #align affine_basis.surjective_coord AffineBasis.surjective_coord diff --git a/Mathbin/LinearAlgebra/AffineSpace/Combination.lean b/Mathbin/LinearAlgebra/AffineSpace/Combination.lean index 93c2c30acf..b3ab7e40e0 100644 --- a/Mathbin/LinearAlgebra/AffineSpace/Combination.lean +++ b/Mathbin/LinearAlgebra/AffineSpace/Combination.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers ! This file was ported from Lean 3 source module linear_algebra.affine_space.combination -! leanprover-community/mathlib commit 87c54600fe3cdc7d32ff5b50873ac724d86aef8d +! leanprover-community/mathlib commit 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -376,6 +376,8 @@ theorem weightedVsub_const_smul (w : ι → k) (p : ι → P) (c : k) : s.weightedVsubOfPoint_const_smul _ _ _ _ #align finset.weighted_vsub_const_smul Finset.weightedVsub_const_smul +variable (k) + /-- A weighted sum of the results of subtracting a default base point from the given points, added to that base point, as an affine map on the weights. This is intended to be used when the sum of the weights @@ -393,10 +395,12 @@ def affineCombination (p : ι → P) : (ι → k) →ᵃ[k] P `weighted_vsub`. -/ @[simp] theorem affineCombination_linear (p : ι → P) : - (s.affineCombination p : (ι → k) →ᵃ[k] P).linear = s.weightedVsub p := + (s.affineCombination k p).linear = s.weightedVsub p := rfl #align finset.affine_combination_linear Finset.affineCombination_linear +variable {k} + /-- Applying `affine_combination` with given weights. This is for the case where a result involving a default base point is OK (for example, when that base point will cancel out later); a more typical use case @@ -405,7 +409,7 @@ point with `affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one` and then using `weighted_vsub_of_point_apply`. -/ theorem affineCombination_apply (w : ι → k) (p : ι → P) : - s.affineCombination p w = + s.affineCombination k p w = s.weightedVsubOfPoint p (Classical.choice S.Nonempty) w +ᵥ Classical.choice S.Nonempty := rfl #align finset.affine_combination_apply Finset.affineCombination_apply @@ -413,14 +417,14 @@ theorem affineCombination_apply (w : ι → k) (p : ι → P) : /-- The value of `affine_combination`, where the given points are equal. -/ @[simp] theorem affineCombination_apply_const (w : ι → k) (p : P) (h : (∑ i in s, w i) = 1) : - s.affineCombination (fun _ => p) w = p := by + s.affineCombination k (fun _ => p) w = p := by rw [affine_combination_apply, s.weighted_vsub_of_point_apply_const, h, one_smul, vsub_vadd] #align finset.affine_combination_apply_const Finset.affineCombination_apply_const /-- `affine_combination` gives equal results for two families of weights and two families of points that are equal on `s`. -/ theorem affineCombination_congr {w₁ w₂ : ι → k} (hw : ∀ i ∈ s, w₁ i = w₂ i) {p₁ p₂ : ι → P} - (hp : ∀ i ∈ s, p₁ i = p₂ i) : s.affineCombination p₁ w₁ = s.affineCombination p₂ w₂ := by + (hp : ∀ i ∈ s, p₁ i = p₂ i) : s.affineCombination k p₁ w₁ = s.affineCombination k p₂ w₂ := by simp_rw [affine_combination_apply, s.weighted_vsub_of_point_congr hw hp] #align finset.affine_combination_congr Finset.affineCombination_congr @@ -428,25 +432,25 @@ theorem affineCombination_congr {w₁ w₂ : ι → k} (hw : ∀ i ∈ s, w₁ i sum of the weights is 1. -/ theorem affineCombination_eq_weightedVsubOfPoint_vadd_of_sum_eq_one (w : ι → k) (p : ι → P) (h : (∑ i in s, w i) = 1) (b : P) : - s.affineCombination p w = s.weightedVsubOfPoint p b w +ᵥ b := + s.affineCombination k p w = s.weightedVsubOfPoint p b w +ᵥ b := s.weightedVsubOfPoint_vadd_eq_of_sum_eq_one w p h _ _ #align finset.affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one Finset.affineCombination_eq_weightedVsubOfPoint_vadd_of_sum_eq_one /-- Adding a `weighted_vsub` to an `affine_combination`. -/ theorem weightedVsub_vadd_affineCombination (w₁ w₂ : ι → k) (p : ι → P) : - s.weightedVsub p w₁ +ᵥ s.affineCombination p w₂ = s.affineCombination p (w₁ + w₂) := by + s.weightedVsub p w₁ +ᵥ s.affineCombination k p w₂ = s.affineCombination k p (w₁ + w₂) := by rw [← vadd_eq_add, AffineMap.map_vadd, affine_combination_linear] #align finset.weighted_vsub_vadd_affine_combination Finset.weightedVsub_vadd_affineCombination /-- Subtracting two `affine_combination`s. -/ theorem affineCombination_vsub (w₁ w₂ : ι → k) (p : ι → P) : - s.affineCombination p w₁ -ᵥ s.affineCombination p w₂ = s.weightedVsub p (w₁ - w₂) := by + s.affineCombination k p w₁ -ᵥ s.affineCombination k p w₂ = s.weightedVsub p (w₁ - w₂) := by rw [← AffineMap.linearMap_vsub, affine_combination_linear, vsub_eq_sub] #align finset.affine_combination_vsub Finset.affineCombination_vsub theorem attach_affineCombination_of_injective [DecidableEq P] (s : Finset P) (w : P → k) (f : s → P) (hf : Function.Injective f) : - s.attach.affineCombination f (w ∘ f) = (image f univ).affineCombination id w := + s.attach.affineCombination k f (w ∘ f) = (image f univ).affineCombination k id w := by simp only [affine_combination, weighted_vsub_of_point_apply, id.def, vadd_right_cancel_iff, Function.comp_apply, AffineMap.coe_mk] @@ -461,7 +465,7 @@ theorem attach_affineCombination_of_injective [DecidableEq P] (s : Finset P) (w #align finset.attach_affine_combination_of_injective Finset.attach_affineCombination_of_injective theorem attach_affineCombination_coe (s : Finset P) (w : P → k) : - s.attach.affineCombination (coe : s → P) (w ∘ coe) = s.affineCombination id w := by + s.attach.affineCombination k (coe : s → P) (w ∘ coe) = s.affineCombination k id w := by classical rw [attach_affine_combination_of_injective s w (coe : s → P) Subtype.coe_injective, univ_eq_attach, attach_image_coe] #align finset.attach_affine_combination_coe Finset.attach_affineCombination_coe @@ -480,7 +484,7 @@ theorem weightedVsub_eq_linear_combination {ι} (s : Finset ι) {w : ι → k} { combinations. -/ @[simp] theorem affineCombination_eq_linear_combination (s : Finset ι) (p : ι → V) (w : ι → k) - (hw : (∑ i in s, w i) = 1) : s.affineCombination p w = ∑ i in s, w i • p i := by + (hw : (∑ i in s, w i) = 1) : s.affineCombination k p w = ∑ i in s, w i • p i := by simp [s.affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one w p hw 0] #align finset.affine_combination_eq_linear_combination Finset.affineCombination_eq_linear_combination @@ -490,7 +494,7 @@ include S and has weight 1 and the other points in the set have weight 0. -/ @[simp] theorem affineCombination_of_eq_one_of_eq_zero (w : ι → k) (p : ι → P) {i : ι} (his : i ∈ s) - (hwi : w i = 1) (hw0 : ∀ i2 ∈ s, i2 ≠ i → w i2 = 0) : s.affineCombination p w = p i := + (hwi : w i = 1) (hw0 : ∀ i2 ∈ s, i2 ≠ i → w i2 = 0) : s.affineCombination k p w = p i := by have h1 : (∑ i in s, w i) = 1 := hwi ▸ sum_eq_single i hw0 fun h => False.elim (h his) rw [s.affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one w p h1 (p i), @@ -506,7 +510,8 @@ theorem affineCombination_of_eq_one_of_eq_zero (w : ι → k) (p : ι → P) {i /-- An affine combination is unaffected by changing the weights to the corresponding indicator function and adding points to the set. -/ theorem affineCombination_indicator_subset (w : ι → k) (p : ι → P) {s₁ s₂ : Finset ι} - (h : s₁ ⊆ s₂) : s₁.affineCombination p w = s₂.affineCombination p (Set.indicator (↑s₁) w) := by + (h : s₁ ⊆ s₂) : s₁.affineCombination k p w = s₂.affineCombination k p (Set.indicator (↑s₁) w) := + by rw [affine_combination_apply, affine_combination_apply, weighted_vsub_of_point_indicator_subset _ _ _ h] #align finset.affine_combination_indicator_subset Finset.affineCombination_indicator_subset @@ -515,14 +520,14 @@ theorem affineCombination_indicator_subset (w : ι → k) (p : ι → P) {s₁ s affine combination with the same points and weights over the original `finset`. -/ theorem affineCombination_map (e : ι₂ ↪ ι) (w : ι → k) (p : ι → P) : - (s₂.map e).affineCombination p w = s₂.affineCombination (p ∘ e) (w ∘ e) := by + (s₂.map e).affineCombination k p w = s₂.affineCombination k (p ∘ e) (w ∘ e) := by simp_rw [affine_combination_apply, weighted_vsub_of_point_map] #align finset.affine_combination_map Finset.affineCombination_map /-- A weighted sum of pairwise subtractions, expressed as a subtraction of two `affine_combination` expressions. -/ theorem sum_smul_vsub_eq_affineCombination_vsub (w : ι → k) (p₁ p₂ : ι → P) : - (∑ i in s, w i • (p₁ i -ᵥ p₂ i)) = s.affineCombination p₁ w -ᵥ s.affineCombination p₂ w := + (∑ i in s, w i • (p₁ i -ᵥ p₂ i)) = s.affineCombination k p₁ w -ᵥ s.affineCombination k p₂ w := by simp_rw [affine_combination_apply, vadd_vsub_vadd_cancel_right] exact s.sum_smul_vsub_eq_weighted_vsub_of_point_sub _ _ _ _ @@ -531,21 +536,21 @@ theorem sum_smul_vsub_eq_affineCombination_vsub (w : ι → k) (p₁ p₂ : ι /-- A weighted sum of pairwise subtractions, where the point on the right is constant and the sum of the weights is 1. -/ theorem sum_smul_vsub_const_eq_affineCombination_vsub (w : ι → k) (p₁ : ι → P) (p₂ : P) - (h : (∑ i in s, w i) = 1) : (∑ i in s, w i • (p₁ i -ᵥ p₂)) = s.affineCombination p₁ w -ᵥ p₂ := + (h : (∑ i in s, w i) = 1) : (∑ i in s, w i • (p₁ i -ᵥ p₂)) = s.affineCombination k p₁ w -ᵥ p₂ := by rw [sum_smul_vsub_eq_affine_combination_vsub, affine_combination_apply_const _ _ _ h] #align finset.sum_smul_vsub_const_eq_affine_combination_vsub Finset.sum_smul_vsub_const_eq_affineCombination_vsub /-- A weighted sum of pairwise subtractions, where the point on the left is constant and the sum of the weights is 1. -/ theorem sum_smul_const_vsub_eq_vsub_affineCombination (w : ι → k) (p₂ : ι → P) (p₁ : P) - (h : (∑ i in s, w i) = 1) : (∑ i in s, w i • (p₁ -ᵥ p₂ i)) = p₁ -ᵥ s.affineCombination p₂ w := + (h : (∑ i in s, w i) = 1) : (∑ i in s, w i • (p₁ -ᵥ p₂ i)) = p₁ -ᵥ s.affineCombination k p₂ w := by rw [sum_smul_vsub_eq_affine_combination_vsub, affine_combination_apply_const _ _ _ h] #align finset.sum_smul_const_vsub_eq_vsub_affine_combination Finset.sum_smul_const_vsub_eq_vsub_affineCombination /-- A weighted sum may be split into a subtraction of affine combinations over two subsets. -/ theorem affineCombination_sdiff_sub [DecidableEq ι] {s₂ : Finset ι} (h : s₂ ⊆ s) (w : ι → k) (p : ι → P) : - (s \ s₂).affineCombination p w -ᵥ s₂.affineCombination p (-w) = s.weightedVsub p w := + (s \ s₂).affineCombination k p w -ᵥ s₂.affineCombination k p (-w) = s.weightedVsub p w := by simp_rw [affine_combination_apply, vadd_vsub_vadd_cancel_right] exact s.weighted_vsub_sdiff_sub h _ _ @@ -555,7 +560,7 @@ theorem affineCombination_sdiff_sub [DecidableEq ι] {s₂ : Finset ι} (h : s the affine combination of the other points with the given weights. -/ theorem affineCombination_eq_of_weightedVsub_eq_zero_of_eq_neg_one {w : ι → k} {p : ι → P} (hw : s.weightedVsub p w = (0 : V)) {i : ι} [DecidablePred (· ≠ i)] (his : i ∈ s) - (hwi : w i = -1) : (s.filterₓ (· ≠ i)).affineCombination p w = p i := by + (hwi : w i = -1) : (s.filterₓ (· ≠ i)).affineCombination k p w = p i := by classical rw [← @vsub_eq_zero_iff_eq V, ← hw, ← s.affine_combination_sdiff_sub (singleton_subset_iff.2 his), sdiff_singleton_eq_erase, ← @@ -569,8 +574,8 @@ theorem affineCombination_eq_of_weightedVsub_eq_zero_of_eq_neg_one {w : ι → k /-- An affine combination over `s.subtype pred` equals one over `s.filter pred`. -/ theorem affineCombination_subtype_eq_filter (w : ι → k) (p : ι → P) (pred : ι → Prop) [DecidablePred pred] : - ((s.Subtype pred).affineCombination (fun i => p i) fun i => w i) = - (s.filterₓ pred).affineCombination p w := + ((s.Subtype pred).affineCombination k (fun i => p i) fun i => w i) = + (s.filterₓ pred).affineCombination k p w := by rw [affine_combination_apply, affine_combination_apply, weighted_vsub_of_point_subtype_eq_filter] #align finset.affine_combination_subtype_eq_filter Finset.affineCombination_subtype_eq_filter @@ -579,7 +584,7 @@ theorem affineCombination_subtype_eq_filter (w : ι → k) (p : ι → P) (pred in `s` not satisfying `pred` are zero. -/ theorem affineCombination_filter_of_ne (w : ι → k) (p : ι → P) {pred : ι → Prop} [DecidablePred pred] (h : ∀ i ∈ s, w i ≠ 0 → pred i) : - (s.filterₓ pred).affineCombination p w = s.affineCombination p w := by + (s.filterₓ pred).affineCombination k p w = s.affineCombination k p w := by rw [affine_combination_apply, affine_combination_apply, s.weighted_vsub_of_point_filter_of_ne _ _ _ h] #align finset.affine_combination_filter_of_ne Finset.affineCombination_filter_of_ne @@ -640,9 +645,9 @@ subset. -/ theorem eq_affineCombination_subset_iff_eq_affineCombination_subtype {p0 : P} {s : Set ι} {p : ι → P} : (∃ (fs : Finset ι)(hfs : ↑fs ⊆ s)(w : ι → k)(hw : (∑ i in fs, w i) = 1), - p0 = fs.affineCombination p w) ↔ + p0 = fs.affineCombination k p w) ↔ ∃ (fs : Finset s)(w : s → k)(hw : (∑ i in fs, w i) = 1), - p0 = fs.affineCombination (fun i : s => p i) w := + p0 = fs.affineCombination k (fun i : s => p i) w := by simp_rw [affine_combination_apply, eq_vadd_iff_vsub_eq] exact eq_weighted_vsub_of_point_subset_iff_eq_weighted_vsub_of_point_subtype @@ -653,7 +658,7 @@ variable {k V} /-- Affine maps commute with affine combinations. -/ theorem map_affineCombination {V₂ P₂ : Type _} [AddCommGroup V₂] [Module k V₂] [affine_space V₂ P₂] (p : ι → P) (w : ι → k) (hw : s.Sum w = 1) (f : P →ᵃ[k] P₂) : - f (s.affineCombination p w) = s.affineCombination (f ∘ p) w := + f (s.affineCombination k p w) = s.affineCombination k (f ∘ p) w := by have b := Classical.choice (inferInstance : affine_space V P).Nonempty have b₂ := Classical.choice (inferInstance : affine_space V₂ P₂).Nonempty @@ -770,7 +775,7 @@ variable (k) /-- An affine combination with `affine_combination_single_weights` gives the specified point. -/ @[simp] theorem affineCombination_affineCombinationSingleWeights [DecidableEq ι] (p : ι → P) {i : ι} - (hi : i ∈ s) : s.affineCombination p (affineCombinationSingleWeights k i) = p i := + (hi : i ∈ s) : s.affineCombination k p (affineCombinationSingleWeights k i) = p i := by refine' s.affine_combination_of_eq_one_of_eq_zero _ _ hi (by simp) _ rintro j - hj @@ -794,7 +799,7 @@ variable {k} @[simp] theorem affineCombination_affineCombinationLineMapWeights [DecidableEq ι] (p : ι → P) {i j : ι} (hi : i ∈ s) (hj : j ∈ s) (c : k) : - s.affineCombination p (affineCombinationLineMapWeights i j c) = + s.affineCombination k p (affineCombinationLineMapWeights i j c) = AffineMap.lineMap (p i) (p j) c := by rw [affine_combination_line_map_weights, ← weighted_vsub_vadd_affine_combination, @@ -862,11 +867,11 @@ include V is intended to be used in the case where the number of points, converted to `k`, is not zero. -/ def centroid (p : ι → P) : P := - s.affineCombination p (s.centroidWeights k) + s.affineCombination k p (s.centroidWeights k) #align finset.centroid Finset.centroid /-- The definition of the centroid. -/ -theorem centroid_def (p : ι → P) : s.centroid k p = s.affineCombination p (s.centroidWeights k) := +theorem centroid_def (p : ι → P) : s.centroid k p = s.affineCombination k p (s.centroidWeights k) := rfl #align finset.centroid_def Finset.centroid_def @@ -973,7 +978,7 @@ include V /-- The centroid as an affine combination over a `fintype`. -/ theorem centroid_eq_affineCombination_fintype [Fintype ι] (p : ι → P) : - s.centroid k p = univ.affineCombination p (s.centroidWeightsIndicator k) := + s.centroid k p = univ.affineCombination k p (s.centroidWeightsIndicator k) := affineCombination_indicator_subset _ _ (subset_univ _) #align finset.centroid_eq_affine_combination_fintype Finset.centroid_eq_affineCombination_fintype @@ -1062,8 +1067,8 @@ theorem weightedVsub_mem_vectorSpan {s : Finset ι} {w : ι → k} (h : (∑ i i `affine_span` of an indexed family, if the underlying ring is nontrivial. -/ theorem affineCombination_mem_affineSpan [Nontrivial k] {s : Finset ι} {w : ι → k} - (h : (∑ i in s, w i) = 1) (p : ι → P) : s.affineCombination p w ∈ affineSpan k (Set.range p) := - by + (h : (∑ i in s, w i) = 1) (p : ι → P) : + s.affineCombination k p w ∈ affineSpan k (Set.range p) := by classical have hnz : (∑ i in s, w i) ≠ 0 := h.symm ▸ one_ne_zero have hn : s.nonempty := Finset.nonempty_of_sum_ne_zero hnz @@ -1071,15 +1076,15 @@ theorem affineCombination_mem_affineSpan [Nontrivial k] {s : Finset ι} {w : ι let w1 : ι → k := Function.update (Function.const ι 0) i1 1 have hw1 : (∑ i in s, w1 i) = 1 := by rw [Finset.sum_update_of_mem hi1, Finset.sum_const_zero, add_zero] - have hw1s : s.affine_combination p w1 = p i1 := + have hw1s : s.affine_combination k p w1 = p i1 := s.affine_combination_of_eq_one_of_eq_zero w1 p hi1 (Function.update_same _ _ _) fun _ _ hne => Function.update_noteq hne _ _ - have hv : s.affine_combination p w -ᵥ p i1 ∈ (affineSpan k (Set.range p)).direction := + have hv : s.affine_combination k p w -ᵥ p i1 ∈ (affineSpan k (Set.range p)).direction := by rw [direction_affineSpan, ← hw1s, Finset.affineCombination_vsub] apply weightedVsub_mem_vectorSpan simp [Pi.sub_apply, h, hw1] - rw [← vsub_vadd (s.affine_combination p w) (p i1)] + rw [← vsub_vadd (s.affine_combination k p w) (p i1)] exact AffineSubspace.vadd_mem_of_mem_direction hv (mem_affineSpan k (Set.mem_range_self _)) #align affine_combination_mem_affine_span affineCombination_mem_affineSpan @@ -1136,7 +1141,7 @@ variable {k} `eq_affine_combination_of_mem_affine_span_of_fintype`. -/ theorem eq_affineCombination_of_mem_affineSpan {p1 : P} {p : ι → P} (h : p1 ∈ affineSpan k (Set.range p)) : - ∃ (s : Finset ι)(w : ι → k)(hw : (∑ i in s, w i) = 1), p1 = s.affineCombination p w := by + ∃ (s : Finset ι)(w : ι → k)(hw : (∑ i in s, w i) = 1), p1 = s.affineCombination k p w := by classical have hn : (affineSpan k (Set.range p) : Set P).Nonempty := ⟨p1, h⟩ rw [affineSpan_nonempty, Set.range_nonempty_iff_nonempty] at hn @@ -1157,7 +1162,7 @@ theorem eq_affineCombination_of_mem_affineSpan {p1 : P} {p : ι → P} let w0 : ι → k := Function.update (Function.const ι 0) i0 1 have hw0 : (∑ i in s', w0 i) = 1 := by rw [Finset.sum_update_of_mem (Finset.mem_insert_self _ _), Finset.sum_const_zero, add_zero] - have hw0s : s'.affine_combination p w0 = p i0 := + have hw0s : s'.affine_combination k p w0 = p i0 := s'.affine_combination_of_eq_one_of_eq_zero w0 p (Finset.mem_insert_self _ _) (Function.update_same _ _ _) fun _ _ hne => Function.update_noteq hne _ _ use s', w0 + w' @@ -1168,7 +1173,7 @@ theorem eq_affineCombination_of_mem_affineSpan {p1 : P} {p : ι → P} theorem eq_affineCombination_of_mem_affineSpan_of_fintype [Fintype ι] {p1 : P} {p : ι → P} (h : p1 ∈ affineSpan k (Set.range p)) : - ∃ (w : ι → k)(hw : (∑ i, w i) = 1), p1 = Finset.univ.affineCombination p w := by + ∃ (w : ι → k)(hw : (∑ i, w i) = 1), p1 = Finset.univ.affineCombination k p w := by classical obtain ⟨s, w, hw, rfl⟩ := eq_affineCombination_of_mem_affineSpan h refine' @@ -1184,7 +1189,7 @@ if it is an `affine_combination` with sum of weights 1, provided the underlying ring is nontrivial. -/ theorem mem_affineSpan_iff_eq_affineCombination [Nontrivial k] {p1 : P} {p : ι → P} : p1 ∈ affineSpan k (Set.range p) ↔ - ∃ (s : Finset ι)(w : ι → k)(hw : (∑ i in s, w i) = 1), p1 = s.affineCombination p w := + ∃ (s : Finset ι)(w : ι → k)(hw : (∑ i in s, w i) = 1), p1 = s.affineCombination k p w := by constructor · exact eq_affineCombination_of_mem_affineSpan diff --git a/Mathbin/LinearAlgebra/AffineSpace/Independent.lean b/Mathbin/LinearAlgebra/AffineSpace/Independent.lean index b0886ffb2b..f3aefedebb 100644 --- a/Mathbin/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathbin/LinearAlgebra/AffineSpace/Independent.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers ! This file was ported from Lean 3 source module linear_algebra.affine_space.independent -! leanprover-community/mathlib commit 09258fb7f75d741b7eda9fa18d5c869e2135d9f1 +! leanprover-community/mathlib commit 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -197,7 +197,7 @@ theorem affineIndependent_iff_indicator_eq_of_affineCombination_eq (p : ι → P ∀ (s1 s2 : Finset ι) (w1 w2 : ι → k), (∑ i in s1, w1 i) = 1 → (∑ i in s2, w2 i) = 1 → - s1.affineCombination p w1 = s2.affineCombination p w2 → + s1.affineCombination k p w1 = s2.affineCombination k p w2 → Set.indicator (↑s1) w1 = Set.indicator (↑s2) w2 := by classical @@ -220,12 +220,12 @@ theorem affineIndependent_iff_indicator_eq_of_affineCombination_eq (p : ι → P let w1 : ι → k := Function.update (Function.const ι 0) i0 1 have hw1 : (∑ i in s, w1 i) = 1 := by rw [Finset.sum_update_of_mem hi0, Finset.sum_const_zero, add_zero] - have hw1s : s.affine_combination p w1 = p i0 := + have hw1s : s.affine_combination k p w1 = p i0 := s.affine_combination_of_eq_one_of_eq_zero w1 p hi0 (Function.update_same _ _ _) fun _ _ hne => Function.update_noteq hne _ _ let w2 := w + w1 have hw2 : (∑ i in s, w2 i) = 1 := by simp [w2, Finset.sum_add_distrib, hw, hw1] - have hw2s : s.affine_combination p w2 = p i0 := by + have hw2s : s.affine_combination k p w2 = p i0 := by simp [w2, ← Finset.weightedVsub_vadd_affineCombination, hs, hw1s] replace ha := ha s s w2 w1 hw2 hw1 (hw1s.symm ▸ hw2s) have hws : w2 i0 - w1 i0 = 0 := by @@ -241,7 +241,7 @@ theorem affineIndependent_iff_eq_of_fintype_affineCombination_eq [Fintype ι] (p ∀ w1 w2 : ι → k, (∑ i, w1 i) = 1 → (∑ i, w2 i) = 1 → - Finset.univ.affineCombination p w1 = Finset.univ.affineCombination p w2 → w1 = w2 := + Finset.univ.affineCombination k p w1 = Finset.univ.affineCombination k p w2 → w1 = w2 := by rw [affineIndependent_iff_indicator_eq_of_affineCombination_eq] constructor @@ -274,7 +274,7 @@ theorem AffineIndependent.units_lineMap {p : ι → P} (hp : AffineIndependent k theorem AffineIndependent.indicator_eq_of_affineCombination_eq {p : ι → P} (ha : AffineIndependent k p) (s₁ s₂ : Finset ι) (w₁ w₂ : ι → k) (hw₁ : (∑ i in s₁, w₁ i) = 1) - (hw₂ : (∑ i in s₂, w₂ i) = 1) (h : s₁.affineCombination p w₁ = s₂.affineCombination p w₂) : + (hw₂ : (∑ i in s₂, w₂ i) = 1) (h : s₁.affineCombination k p w₁ = s₂.affineCombination k p w₂) : Set.indicator (↑s₁) w₁ = Set.indicator (↑s₂) w₂ := (affineIndependent_iff_indicator_eq_of_affineCombination_eq k p).1 ha s₁ s₂ w₁ w₂ hw₁ hw₂ h #align affine_independent.indicator_eq_of_affine_combination_eq AffineIndependent.indicator_eq_of_affineCombination_eq @@ -521,7 +521,7 @@ theorem weightedVsub_mem_vectorSpan_pair {p : ι → P} (h : AffineIndependent k {s : Finset ι} (hw : (∑ i in s, w i) = 0) (hw₁ : (∑ i in s, w₁ i) = 1) (hw₂ : (∑ i in s, w₂ i) = 1) : s.weightedVsub p w ∈ - vectorSpan k ({s.affineCombination p w₁, s.affineCombination p w₂} : Set P) ↔ + vectorSpan k ({s.affineCombination k p w₁, s.affineCombination k p w₂} : Set P) ↔ ∃ r : k, ∀ i ∈ s, w i = r * (w₁ i - w₂ i) := by rw [mem_vectorSpan_pair] @@ -551,10 +551,10 @@ two points. -/ theorem affineCombination_mem_affineSpan_pair {p : ι → P} (h : AffineIndependent k p) {w w₁ w₂ : ι → k} {s : Finset ι} (hw : (∑ i in s, w i) = 1) (hw₁ : (∑ i in s, w₁ i) = 1) (hw₂ : (∑ i in s, w₂ i) = 1) : - s.affineCombination p w ∈ line[k, s.affineCombination p w₁, s.affineCombination p w₂] ↔ + s.affineCombination k p w ∈ line[k, s.affineCombination k p w₁, s.affineCombination k p w₂] ↔ ∃ r : k, ∀ i ∈ s, w i = r * (w₂ i - w₁ i) + w₁ i := by - rw [← vsub_vadd (s.affine_combination p w) (s.affine_combination p w₁), + rw [← vsub_vadd (s.affine_combination k p w) (s.affine_combination k p w₁), AffineSubspace.vadd_mem_iff_mem_direction _ (left_mem_affineSpan_pair _ _ _), direction_affineSpan, s.affine_combination_vsub, Set.pair_comm, weightedVsub_mem_vectorSpan_pair h _ hw₂ hw₁] @@ -771,7 +771,8 @@ sign. -/ theorem sign_eq_of_affineCombination_mem_affineSpan_pair {p : ι → P} (h : AffineIndependent k p) {w w₁ w₂ : ι → k} {s : Finset ι} (hw : (∑ i in s, w i) = 1) (hw₁ : (∑ i in s, w₁ i) = 1) (hw₂ : (∑ i in s, w₂ i) = 1) - (hs : s.affineCombination p w ∈ line[k, s.affineCombination p w₁, s.affineCombination p w₂]) + (hs : + s.affineCombination k p w ∈ line[k, s.affineCombination k p w₁, s.affineCombination k p w₂]) {i j : ι} (hi : i ∈ s) (hj : j ∈ s) (hi0 : w₁ i = 0) (hj0 : w₁ j = 0) (hij : SignType.sign (w₂ i) = SignType.sign (w₂ j)) : SignType.sign (w i) = SignType.sign (w j) := @@ -790,7 +791,7 @@ theorem sign_eq_of_affineCombination_mem_affineSpan_single_lineMap {p : ι → P (h : AffineIndependent k p) {w : ι → k} {s : Finset ι} (hw : (∑ i in s, w i) = 1) {i₁ i₂ i₃ : ι} (h₁ : i₁ ∈ s) (h₂ : i₂ ∈ s) (h₃ : i₃ ∈ s) (h₁₂ : i₁ ≠ i₂) (h₁₃ : i₁ ≠ i₃) (h₂₃ : i₂ ≠ i₃) {c : k} (hc0 : 0 < c) (hc1 : c < 1) - (hs : s.affineCombination p w ∈ line[k, p i₁, AffineMap.lineMap (p i₂) (p i₃) c]) : + (hs : s.affineCombination k p w ∈ line[k, p i₁, AffineMap.lineMap (p i₂) (p i₃) c]) : SignType.sign (w i₂) = SignType.sign (w i₃) := by classical rw [← s.affine_combination_affine_combination_single_weights k p h₁, ← diff --git a/Mathbin/LinearAlgebra/AffineSpace/Matrix.lean b/Mathbin/LinearAlgebra/AffineSpace/Matrix.lean index bb5e6393b3..92967970f1 100644 --- a/Mathbin/LinearAlgebra/AffineSpace/Matrix.lean +++ b/Mathbin/LinearAlgebra/AffineSpace/Matrix.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash ! This file was ported from Lean 3 source module linear_algebra.affine_space.matrix -! leanprover-community/mathlib commit 2f4cdce0c2f2f3b8cd58f05d556d03b468e1eb2e +! leanprover-community/mathlib commit 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -99,7 +99,7 @@ theorem affineSpan_eq_top_of_toMatrix_left_inv [DecidableEq ι] [Nontrivial k] ( _ = ∑ l, (A ⬝ b.to_matrix p) i l := rfl _ = 1 := by simp [hA, Matrix.one_apply, Finset.filter_eq] - have hbi : b i = finset.univ.affine_combination p (A i) := + have hbi : b i = finset.univ.affine_combination k p (A i) := by apply b.ext_elem intro j diff --git a/Mathbin/MeasureTheory/Group/FundamentalDomain.lean b/Mathbin/MeasureTheory/Group/FundamentalDomain.lean index ec2344fd98..5390653d8d 100644 --- a/Mathbin/MeasureTheory/Group/FundamentalDomain.lean +++ b/Mathbin/MeasureTheory/Group/FundamentalDomain.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov ! This file was ported from Lean 3 source module measure_theory.group.fundamental_domain -! leanprover-community/mathlib commit f2ce6086713c78a7f880485f7917ea547a215982 +! leanprover-community/mathlib commit eb810cf549db839dadf13353dbe69bac55acdbbc ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -254,21 +254,35 @@ theorem lintegral_eq_tsum (h : IsFundamentalDomain G s μ) (f : α → ℝ≥0 #align measure_theory.is_add_fundamental_domain.lintegral_eq_tsum MeasureTheory.IsAddFundamentalDomain.lintegral_eq_tsum @[to_additive] -theorem set_lintegral_eq_tsum' (h : IsFundamentalDomain G s μ) (f : α → ℝ≥0∞) (t : Set α) : +theorem lintegral_eq_tsum' (h : IsFundamentalDomain G s μ) (f : α → ℝ≥0∞) : + (∫⁻ x, f x ∂μ) = ∑' g : G, ∫⁻ x in s, f (g⁻¹ • x) ∂μ := + calc + (∫⁻ x, f x ∂μ) = ∑' g : G, ∫⁻ x in g • s, f x ∂μ := h.lintegral_eq_tsum f + _ = ∑' g : G, ∫⁻ x in g⁻¹ • s, f x ∂μ := ((Equiv.inv G).tsum_eq _).symm + _ = ∑' g : G, ∫⁻ x in s, f (g⁻¹ • x) ∂μ := + tsum_congr fun g => + ((measurePreservingSmul g⁻¹ μ).set_lintegral_comp_emb (measurableEmbedding_const_smul _) _ + _).symm + +#align measure_theory.is_fundamental_domain.lintegral_eq_tsum' MeasureTheory.IsFundamentalDomain.lintegral_eq_tsum' +#align measure_theory.is_add_fundamental_domain.lintegral_eq_tsum' MeasureTheory.IsAddFundamentalDomain.lintegral_eq_tsum' + +@[to_additive] +theorem set_lintegral_eq_tsum (h : IsFundamentalDomain G s μ) (f : α → ℝ≥0∞) (t : Set α) : (∫⁻ x in t, f x ∂μ) = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ := calc (∫⁻ x in t, f x ∂μ) = ∑' g : G, ∫⁻ x in g • s, f x ∂μ.restrict t := h.lintegral_eq_tsum_of_ac restrict_le_self.AbsolutelyContinuous _ _ = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ := by simp only [h.restrict_restrict, inter_comm] -#align measure_theory.is_fundamental_domain.set_lintegral_eq_tsum' MeasureTheory.IsFundamentalDomain.set_lintegral_eq_tsum' -#align measure_theory.is_add_fundamental_domain.set_lintegral_eq_tsum' MeasureTheory.IsAddFundamentalDomain.set_lintegral_eq_tsum' +#align measure_theory.is_fundamental_domain.set_lintegral_eq_tsum MeasureTheory.IsFundamentalDomain.set_lintegral_eq_tsum +#align measure_theory.is_add_fundamental_domain.set_lintegral_eq_tsum MeasureTheory.IsAddFundamentalDomain.set_lintegral_eq_tsum @[to_additive] -theorem set_lintegral_eq_tsum (h : IsFundamentalDomain G s μ) (f : α → ℝ≥0∞) (t : Set α) : +theorem set_lintegral_eq_tsum' (h : IsFundamentalDomain G s μ) (f : α → ℝ≥0∞) (t : Set α) : (∫⁻ x in t, f x ∂μ) = ∑' g : G, ∫⁻ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := calc - (∫⁻ x in t, f x ∂μ) = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ := h.set_lintegral_eq_tsum' f t + (∫⁻ x in t, f x ∂μ) = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ := h.set_lintegral_eq_tsum f t _ = ∑' g : G, ∫⁻ x in t ∩ g⁻¹ • s, f x ∂μ := ((Equiv.inv G).tsum_eq _).symm _ = ∑' g : G, ∫⁻ x in g⁻¹ • (g • t ∩ s), f x ∂μ := by simp only [smul_set_inter, inv_smul_smul] _ = ∑' g : G, ∫⁻ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := @@ -276,8 +290,8 @@ theorem set_lintegral_eq_tsum (h : IsFundamentalDomain G s μ) (f : α → ℝ ((measurePreservingSmul g⁻¹ μ).set_lintegral_comp_emb (measurableEmbedding_const_smul _) _ _).symm -#align measure_theory.is_fundamental_domain.set_lintegral_eq_tsum MeasureTheory.IsFundamentalDomain.set_lintegral_eq_tsum -#align measure_theory.is_add_fundamental_domain.set_lintegral_eq_tsum MeasureTheory.IsAddFundamentalDomain.set_lintegral_eq_tsum +#align measure_theory.is_fundamental_domain.set_lintegral_eq_tsum' MeasureTheory.IsFundamentalDomain.set_lintegral_eq_tsum' +#align measure_theory.is_add_fundamental_domain.set_lintegral_eq_tsum' MeasureTheory.IsAddFundamentalDomain.set_lintegral_eq_tsum' @[to_additive] theorem measure_eq_tsum_of_ac (h : IsFundamentalDomain G s μ) (hν : ν ≪ μ) (t : Set α) : @@ -300,7 +314,7 @@ theorem measure_eq_tsum' (h : IsFundamentalDomain G s μ) (t : Set α) : @[to_additive] theorem measure_eq_tsum (h : IsFundamentalDomain G s μ) (t : Set α) : μ t = ∑' g : G, μ (g • t ∩ s) := by - simpa only [set_lintegral_one] using h.set_lintegral_eq_tsum (fun _ => 1) t + simpa only [set_lintegral_one] using h.set_lintegral_eq_tsum' (fun _ => 1) t #align measure_theory.is_fundamental_domain.measure_eq_tsum MeasureTheory.IsFundamentalDomain.measure_eq_tsum #align measure_theory.is_add_fundamental_domain.measure_eq_tsum MeasureTheory.IsAddFundamentalDomain.measure_eq_tsum @@ -332,9 +346,9 @@ protected theorem set_lintegral_eq (hs : IsFundamentalDomain G s μ) (ht : IsFun (f : α → ℝ≥0∞) (hf : ∀ (g : G) (x), f (g • x) = f x) : (∫⁻ x in s, f x ∂μ) = ∫⁻ x in t, f x ∂μ := calc - (∫⁻ x in s, f x ∂μ) = ∑' g : G, ∫⁻ x in s ∩ g • t, f x ∂μ := ht.set_lintegral_eq_tsum' _ _ + (∫⁻ x in s, f x ∂μ) = ∑' g : G, ∫⁻ x in s ∩ g • t, f x ∂μ := ht.set_lintegral_eq_tsum _ _ _ = ∑' g : G, ∫⁻ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := by simp only [hf, inter_comm] - _ = ∫⁻ x in t, f x ∂μ := (hs.set_lintegral_eq_tsum _ _).symm + _ = ∫⁻ x in t, f x ∂μ := (hs.set_lintegral_eq_tsum' _ _).symm #align measure_theory.is_fundamental_domain.set_lintegral_eq MeasureTheory.IsFundamentalDomain.set_lintegral_eq #align measure_theory.is_add_fundamental_domain.set_lintegral_eq MeasureTheory.IsAddFundamentalDomain.set_lintegral_eq @@ -412,31 +426,73 @@ protected theorem integrableOn_iff (hs : IsFundamentalDomain G s μ) (ht : IsFun variable [NormedSpace ℝ E] [CompleteSpace E] +@[to_additive] +theorem integral_eq_tsum_of_ac (h : IsFundamentalDomain G s μ) (hν : ν ≪ μ) (f : α → E) + (hf : Integrable f ν) : (∫ x, f x ∂ν) = ∑' g : G, ∫ x in g • s, f x ∂ν := + by + rw [← MeasureTheory.integral_sum_measure, h.sum_restrict_of_ac hν] + rw [h.sum_restrict_of_ac hν] + -- Weirdly, these rewrites seem not to be combinable + exact hf +#align measure_theory.is_fundamental_domain.integral_eq_tsum_of_ac MeasureTheory.IsFundamentalDomain.integral_eq_tsum_of_ac +#align measure_theory.is_add_fundamental_domain.integral_eq_tsum_of_ac MeasureTheory.IsAddFundamentalDomain.integral_eq_tsum_of_ac + +@[to_additive] +theorem integral_eq_tsum (h : IsFundamentalDomain G s μ) (f : α → E) (hf : Integrable f μ) : + (∫ x, f x ∂μ) = ∑' g : G, ∫ x in g • s, f x ∂μ := + integral_eq_tsum_of_ac h (by rfl) f hf +#align measure_theory.is_fundamental_domain.integral_eq_tsum MeasureTheory.IsFundamentalDomain.integral_eq_tsum +#align measure_theory.is_add_fundamental_domain.integral_eq_tsum MeasureTheory.IsAddFundamentalDomain.integral_eq_tsum + +@[to_additive] +theorem integral_eq_tsum' (h : IsFundamentalDomain G s μ) (f : α → E) (hf : Integrable f μ) : + (∫ x, f x ∂μ) = ∑' g : G, ∫ x in s, f (g⁻¹ • x) ∂μ := + calc + (∫ x, f x ∂μ) = ∑' g : G, ∫ x in g • s, f x ∂μ := h.integral_eq_tsum f hf + _ = ∑' g : G, ∫ x in g⁻¹ • s, f x ∂μ := ((Equiv.inv G).tsum_eq _).symm + _ = ∑' g : G, ∫ x in s, f (g⁻¹ • x) ∂μ := + tsum_congr fun g => + (measurePreservingSmul g⁻¹ μ).set_integral_image_emb (measurableEmbedding_const_smul _) _ _ + +#align measure_theory.is_fundamental_domain.integral_eq_tsum' MeasureTheory.IsFundamentalDomain.integral_eq_tsum' +#align measure_theory.is_add_fundamental_domain.integral_eq_tsum' MeasureTheory.IsAddFundamentalDomain.integral_eq_tsum' + +@[to_additive] +theorem set_integral_eq_tsum (h : IsFundamentalDomain G s μ) {f : α → E} {t : Set α} + (hf : IntegrableOn f t μ) : (∫ x in t, f x ∂μ) = ∑' g : G, ∫ x in t ∩ g • s, f x ∂μ := + calc + (∫ x in t, f x ∂μ) = ∑' g : G, ∫ x in g • s, f x ∂μ.restrict t := + h.integral_eq_tsum_of_ac restrict_le_self.AbsolutelyContinuous f hf + _ = ∑' g : G, ∫ x in t ∩ g • s, f x ∂μ := by + simp only [h.restrict_restrict, measure_smul, inter_comm] + +#align measure_theory.is_fundamental_domain.set_integral_eq_tsum MeasureTheory.IsFundamentalDomain.set_integral_eq_tsum +#align measure_theory.is_add_fundamental_domain.set_integral_eq_tsum MeasureTheory.IsAddFundamentalDomain.set_integral_eq_tsum + +@[to_additive] +theorem set_integral_eq_tsum' (h : IsFundamentalDomain G s μ) {f : α → E} {t : Set α} + (hf : IntegrableOn f t μ) : (∫ x in t, f x ∂μ) = ∑' g : G, ∫ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := + calc + (∫ x in t, f x ∂μ) = ∑' g : G, ∫ x in t ∩ g • s, f x ∂μ := h.set_integral_eq_tsum hf + _ = ∑' g : G, ∫ x in t ∩ g⁻¹ • s, f x ∂μ := ((Equiv.inv G).tsum_eq _).symm + _ = ∑' g : G, ∫ x in g⁻¹ • (g • t ∩ s), f x ∂μ := by simp only [smul_set_inter, inv_smul_smul] + _ = ∑' g : G, ∫ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := + tsum_congr fun g => + (measurePreservingSmul g⁻¹ μ).set_integral_image_emb (measurableEmbedding_const_smul _) _ _ + +#align measure_theory.is_fundamental_domain.set_integral_eq_tsum' MeasureTheory.IsFundamentalDomain.set_integral_eq_tsum' +#align measure_theory.is_add_fundamental_domain.set_integral_eq_tsum' MeasureTheory.IsAddFundamentalDomain.set_integral_eq_tsum' + @[to_additive] protected theorem set_integral_eq (hs : IsFundamentalDomain G s μ) (ht : IsFundamentalDomain G t μ) {f : α → E} (hf : ∀ (g : G) (x), f (g • x) = f x) : (∫ x in s, f x ∂μ) = ∫ x in t, f x ∂μ := by by_cases hfs : integrable_on f s μ · have hft : integrable_on f t μ := by rwa [ht.integrable_on_iff hs hf] - have hac : ∀ {u}, μ.restrict u ≪ μ := fun u => restrict_le_self.absolutely_continuous calc - (∫ x in s, f x ∂μ) = ∫ x in ⋃ g : G, g • t, f x ∂μ.restrict s := by - rw [restrict_congr_set (hac ht.Union_smul_ae_eq), restrict_univ] - _ = ∑' g : G, ∫ x in g • t, f x ∂μ.restrict s := - (integral_Union_ae (fun g => (ht.null_measurable_set_smul g).monoAc hac) - (ht.pairwise_ae_disjoint_of_ac hac) hfs.integrable.integrable_on) - _ = ∑' g : G, ∫ x in s ∩ g • t, f x ∂μ := by simp only [ht.restrict_restrict, inter_comm] - _ = ∑' g : G, ∫ x in s ∩ g⁻¹ • t, f x ∂μ := ((Equiv.inv G).tsum_eq _).symm - _ = ∑' g : G, ∫ x in g⁻¹ • (g • s ∩ t), f x ∂μ := by simp only [smul_set_inter, inv_smul_smul] - _ = ∑' g : G, ∫ x in g • s ∩ t, f (g⁻¹ • x) ∂μ := - (tsum_congr fun g => - (measure_preserving_smul g⁻¹ μ).set_integral_image_emb (measurableEmbedding_const_smul _) - _ _) - _ = ∑' g : G, ∫ x in g • s, f x ∂μ.restrict t := by simp only [hf, hs.restrict_restrict] - _ = ∫ x in ⋃ g : G, g • s, f x ∂μ.restrict t := - (integral_Union_ae (fun g => (hs.null_measurable_set_smul g).monoAc hac) - (hs.ae_disjoint.mono fun i j h => hac h) hft.integrable.integrable_on).symm - _ = ∫ x in t, f x ∂μ := by rw [restrict_congr_set (hac hs.Union_smul_ae_eq), restrict_univ] + (∫ x in s, f x ∂μ) = ∑' g : G, ∫ x in s ∩ g • t, f x ∂μ := ht.set_integral_eq_tsum hfs + _ = ∑' g : G, ∫ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := by simp only [hf, inter_comm] + _ = ∫ x in t, f x ∂μ := (hs.set_integral_eq_tsum' hft).symm · rw [integral_undef hfs, integral_undef] rwa [hs.integrable_on_iff ht hf] at hfs diff --git a/Mathbin/Order/Category/BddDistLat.lean b/Mathbin/Order/Category/BddDistLat.lean new file mode 100644 index 0000000000..1dcdb36ca4 --- /dev/null +++ b/Mathbin/Order/Category/BddDistLat.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies + +! This file was ported from Lean 3 source module order.category.BddDistLat +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Order.Category.BddLat +import Mathbin.Order.Category.DistLat + +/-! +# The category of bounded distributive lattices + +This defines `BddDistLat`, the category of bounded distributive lattices. + +Note that this category is sometimes called [`DistLat`](https://ncatlab.org/nlab/show/DistLat) when +being a lattice is understood to entail having a bottom and a top element. +-/ + + +universe u + +open CategoryTheory + +/-- The category of bounded distributive lattices with bounded lattice morphisms. -/ +structure BddDistLat where + toDistLat : DistLat + [isBoundedOrder : BoundedOrder to_DistLat] +#align BddDistLat BddDistLat + +namespace BddDistLat + +instance : CoeSort BddDistLat (Type _) := + ⟨fun X => X.toDistLat⟩ + +instance (X : BddDistLat) : DistribLattice X := + X.toDistLat.str + +attribute [instance] BddDistLat.isBoundedOrder + +/-- Construct a bundled `BddDistLat` from a `bounded_order` `distrib_lattice`. -/ +def of (α : Type _) [DistribLattice α] [BoundedOrder α] : BddDistLat := + ⟨⟨α⟩⟩ +#align BddDistLat.of BddDistLat.of + +@[simp] +theorem coe_of (α : Type _) [DistribLattice α] [BoundedOrder α] : ↥(of α) = α := + rfl +#align BddDistLat.coe_of BddDistLat.coe_of + +instance : Inhabited BddDistLat := + ⟨of PUnit⟩ + +/-- Turn a `BddDistLat` into a `BddLat` by forgetting it is distributive. -/ +def toBddLat (X : BddDistLat) : BddLat := + BddLat.of X +#align BddDistLat.to_BddLat BddDistLat.toBddLat + +@[simp] +theorem coe_toBddLat (X : BddDistLat) : ↥X.toBddLat = ↥X := + rfl +#align BddDistLat.coe_to_BddLat BddDistLat.coe_toBddLat + +instance : LargeCategory.{u} BddDistLat := + InducedCategory.category toBddLat + +instance : ConcreteCategory BddDistLat := + InducedCategory.concreteCategory toBddLat + +instance hasForgetToDistLat : HasForget₂ BddDistLat DistLat + where forget₂ := + { obj := fun X => ⟨X⟩ + map := fun X Y => BoundedLatticeHom.toLatticeHom } +#align BddDistLat.has_forget_to_DistLat BddDistLat.hasForgetToDistLat + +instance hasForgetToBddLat : HasForget₂ BddDistLat BddLat := + InducedCategory.hasForget₂ toBddLat +#align BddDistLat.has_forget_to_BddLat BddDistLat.hasForgetToBddLat + +theorem forget_bddLat_lat_eq_forget_distLat_lat : + forget₂ BddDistLat BddLat ⋙ forget₂ BddLat Lat = + forget₂ BddDistLat DistLat ⋙ forget₂ DistLat Lat := + rfl +#align BddDistLat.forget_BddLat_Lat_eq_forget_DistLat_Lat BddDistLat.forget_bddLat_lat_eq_forget_distLat_lat + +/-- Constructs an equivalence between bounded distributive lattices from an order isomorphism +between them. -/ +@[simps] +def Iso.mk {α β : BddDistLat.{u}} (e : α ≃o β) : α ≅ β + where + Hom := (e : BoundedLatticeHom α β) + inv := (e.symm : BoundedLatticeHom β α) + hom_inv_id' := by + ext + exact e.symm_apply_apply _ + inv_hom_id' := by + ext + exact e.apply_symm_apply _ +#align BddDistLat.iso.mk BddDistLat.Iso.mk + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : BddDistLat ⥤ BddDistLat where + obj X := of Xᵒᵈ + map X Y := BoundedLatticeHom.dual +#align BddDistLat.dual BddDistLat.dual + +/-- The equivalence between `BddDistLat` and itself induced by `order_dual` both ways. -/ +@[simps Functor inverse] +def dualEquiv : BddDistLat ≌ BddDistLat := + 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 BddDistLat.dual_equiv BddDistLat.dualEquiv + +end BddDistLat + +theorem bddDistLat_dual_comp_forget_to_distLat : + BddDistLat.dual ⋙ forget₂ BddDistLat DistLat = forget₂ BddDistLat DistLat ⋙ DistLat.dual := + rfl +#align BddDistLat_dual_comp_forget_to_DistLat bddDistLat_dual_comp_forget_to_distLat + diff --git a/Mathbin/Order/Category/BddLat.lean b/Mathbin/Order/Category/BddLat.lean new file mode 100644 index 0000000000..ac71e6226e --- /dev/null +++ b/Mathbin/Order/Category/BddLat.lean @@ -0,0 +1,184 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies + +! This file was ported from Lean 3 source module order.category.BddLat +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Order.Category.BddOrd +import Mathbin.Order.Category.Lat +import Mathbin.Order.Category.Semilat + +/-! +# The category of bounded lattices + +This file defines `BddLat`, the category of bounded lattices. + +In literature, this is sometimes called `Lat`, the category of lattices, because being a lattice is +understood to entail having a bottom and a top element. +-/ + + +universe u + +open CategoryTheory + +/-- The category of bounded lattices with bounded lattice morphisms. -/ +structure BddLat where + toLat : Lat + [isBoundedOrder : BoundedOrder to_Lat] +#align BddLat BddLat + +namespace BddLat + +instance : CoeSort BddLat (Type _) := + ⟨fun X => X.toLat⟩ + +instance (X : BddLat) : Lattice X := + X.toLat.str + +attribute [instance] BddLat.isBoundedOrder + +/-- Construct a bundled `BddLat` from `lattice` + `bounded_order`. -/ +def of (α : Type _) [Lattice α] [BoundedOrder α] : BddLat := + ⟨⟨α⟩⟩ +#align BddLat.of BddLat.of + +@[simp] +theorem coe_of (α : Type _) [Lattice α] [BoundedOrder α] : ↥(of α) = α := + rfl +#align BddLat.coe_of BddLat.coe_of + +instance : Inhabited BddLat := + ⟨of PUnit⟩ + +instance : LargeCategory.{u} BddLat + where + Hom X Y := BoundedLatticeHom X Y + id X := BoundedLatticeHom.id X + comp X Y Z f g := g.comp f + id_comp' X Y := BoundedLatticeHom.comp_id + comp_id' X Y := BoundedLatticeHom.id_comp + assoc' W X Y Z _ _ _ := BoundedLatticeHom.comp_assoc _ _ _ + +instance : ConcreteCategory BddLat + where + forget := ⟨coeSort, fun X Y => coeFn, fun X => rfl, fun X Y Z f g => rfl⟩ + forget_faithful := ⟨fun X Y => by convert FunLike.coe_injective⟩ + +instance hasForgetToBddOrd : HasForget₂ BddLat BddOrd + where forget₂ := + { obj := fun X => BddOrd.of X + map := fun X Y => BoundedLatticeHom.toBoundedOrderHom } +#align BddLat.has_forget_to_BddOrd BddLat.hasForgetToBddOrd + +instance hasForgetToLat : HasForget₂ BddLat Lat + where forget₂ := + { obj := fun X => ⟨X⟩ + map := fun X Y => BoundedLatticeHom.toLatticeHom } +#align BddLat.has_forget_to_Lat BddLat.hasForgetToLat + +instance hasForgetToSemilatSup : HasForget₂ BddLat SemilatSup + where forget₂ := + { obj := fun X => ⟨X⟩ + map := fun X Y => BoundedLatticeHom.toSupBotHom } +#align BddLat.has_forget_to_SemilatSup BddLat.hasForgetToSemilatSup + +instance hasForgetToSemilatInf : HasForget₂ BddLat SemilatInf + where forget₂ := + { obj := fun X => ⟨X⟩ + map := fun X Y => BoundedLatticeHom.toInfTopHom } +#align BddLat.has_forget_to_SemilatInf BddLat.hasForgetToSemilatInf + +@[simp] +theorem coe_forget_to_bddOrd (X : BddLat) : ↥((forget₂ BddLat BddOrd).obj X) = ↥X := + rfl +#align BddLat.coe_forget_to_BddOrd BddLat.coe_forget_to_bddOrd + +@[simp] +theorem coe_forget_to_lat (X : BddLat) : ↥((forget₂ BddLat Lat).obj X) = ↥X := + rfl +#align BddLat.coe_forget_to_Lat BddLat.coe_forget_to_lat + +@[simp] +theorem coe_forget_to_semilatSup (X : BddLat) : ↥((forget₂ BddLat SemilatSup).obj X) = ↥X := + rfl +#align BddLat.coe_forget_to_SemilatSup BddLat.coe_forget_to_semilatSup + +@[simp] +theorem coe_forget_to_semilatInf (X : BddLat) : ↥((forget₂ BddLat SemilatInf).obj X) = ↥X := + rfl +#align BddLat.coe_forget_to_SemilatInf BddLat.coe_forget_to_semilatInf + +theorem forget_lat_partOrd_eq_forget_bddOrd_partOrd : + forget₂ BddLat Lat ⋙ forget₂ Lat PartOrd = forget₂ BddLat BddOrd ⋙ forget₂ BddOrd PartOrd := + rfl +#align BddLat.forget_Lat_PartOrd_eq_forget_BddOrd_PartOrd BddLat.forget_lat_partOrd_eq_forget_bddOrd_partOrd + +theorem forget_semilatSup_partOrd_eq_forget_bddOrd_partOrd : + forget₂ BddLat SemilatSup ⋙ forget₂ SemilatSup PartOrd = + forget₂ BddLat BddOrd ⋙ forget₂ BddOrd PartOrd := + rfl +#align BddLat.forget_SemilatSup_PartOrd_eq_forget_BddOrd_PartOrd BddLat.forget_semilatSup_partOrd_eq_forget_bddOrd_partOrd + +theorem forget_semilatInf_partOrd_eq_forget_bddOrd_partOrd : + forget₂ BddLat SemilatInf ⋙ forget₂ SemilatInf PartOrd = + forget₂ BddLat BddOrd ⋙ forget₂ BddOrd PartOrd := + rfl +#align BddLat.forget_SemilatInf_PartOrd_eq_forget_BddOrd_PartOrd BddLat.forget_semilatInf_partOrd_eq_forget_bddOrd_partOrd + +/-- Constructs an equivalence between bounded lattices from an order isomorphism +between them. -/ +@[simps] +def Iso.mk {α β : BddLat.{u}} (e : α ≃o β) : α ≅ β + where + Hom := e + inv := e.symm + hom_inv_id' := by + ext + exact e.symm_apply_apply _ + inv_hom_id' := by + ext + exact e.apply_symm_apply _ +#align BddLat.iso.mk BddLat.Iso.mk + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : BddLat ⥤ BddLat where + obj X := of Xᵒᵈ + map X Y := BoundedLatticeHom.dual +#align BddLat.dual BddLat.dual + +/-- The equivalence between `BddLat` and itself induced by `order_dual` both ways. -/ +@[simps Functor inverse] +def dualEquiv : BddLat ≌ BddLat := + 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 BddLat.dual_equiv BddLat.dualEquiv + +end BddLat + +theorem bddLat_dual_comp_forget_to_bddOrd : + BddLat.dual ⋙ forget₂ BddLat BddOrd = forget₂ BddLat BddOrd ⋙ BddOrd.dual := + rfl +#align BddLat_dual_comp_forget_to_BddOrd bddLat_dual_comp_forget_to_bddOrd + +theorem bddLat_dual_comp_forget_to_lat : + BddLat.dual ⋙ forget₂ BddLat Lat = forget₂ BddLat Lat ⋙ Lat.dual := + rfl +#align BddLat_dual_comp_forget_to_Lat bddLat_dual_comp_forget_to_lat + +theorem bddLat_dual_comp_forget_to_semilatSup : + BddLat.dual ⋙ forget₂ BddLat SemilatSup = forget₂ BddLat SemilatInf ⋙ SemilatInf.dual := + rfl +#align BddLat_dual_comp_forget_to_SemilatSup bddLat_dual_comp_forget_to_semilatSup + +theorem bddLat_dual_comp_forget_to_semilatInf : + BddLat.dual ⋙ forget₂ BddLat SemilatInf = forget₂ BddLat SemilatSup ⋙ SemilatSup.dual := + rfl +#align BddLat_dual_comp_forget_to_SemilatInf bddLat_dual_comp_forget_to_semilatInf + diff --git a/Mathbin/Order/Category/BddOrd.lean b/Mathbin/Order/Category/BddOrd.lean new file mode 100644 index 0000000000..811a2ea227 --- /dev/null +++ b/Mathbin/Order/Category/BddOrd.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies + +! This file was ported from Lean 3 source module order.category.BddOrd +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Category.Bipointed +import Mathbin.Order.Category.PartOrd +import Mathbin.Order.Hom.Bounded + +/-! +# The category of bounded orders + +This defines `BddOrd`, the category of bounded orders. +-/ + + +universe u v + +open CategoryTheory + +/-- The category of bounded orders with monotone functions. -/ +structure BddOrd where + toPartOrd : PartOrd + [isBoundedOrder : BoundedOrder to_PartOrd] +#align BddOrd BddOrd + +namespace BddOrd + +instance : CoeSort BddOrd (Type _) := + InducedCategory.hasCoeToSort toPartOrd + +instance (X : BddOrd) : PartialOrder X := + X.toPartOrd.str + +attribute [instance] BddOrd.isBoundedOrder + +/-- Construct a bundled `BddOrd` from a `fintype` `partial_order`. -/ +def of (α : Type _) [PartialOrder α] [BoundedOrder α] : BddOrd := + ⟨⟨α⟩⟩ +#align BddOrd.of BddOrd.of + +@[simp] +theorem coe_of (α : Type _) [PartialOrder α] [BoundedOrder α] : ↥(of α) = α := + rfl +#align BddOrd.coe_of BddOrd.coe_of + +instance : Inhabited BddOrd := + ⟨of PUnit⟩ + +instance largeCategory : LargeCategory.{u} BddOrd + where + Hom X Y := BoundedOrderHom X Y + id X := BoundedOrderHom.id X + comp X Y Z f g := g.comp f + id_comp' X Y := BoundedOrderHom.comp_id + comp_id' X Y := BoundedOrderHom.id_comp + assoc' W X Y Z _ _ _ := BoundedOrderHom.comp_assoc _ _ _ +#align BddOrd.large_category BddOrd.largeCategory + +instance concreteCategory : ConcreteCategory BddOrd + where + forget := ⟨coeSort, fun X Y => coeFn, fun X => rfl, fun X Y Z f g => rfl⟩ + forget_faithful := ⟨fun X Y => by convert FunLike.coe_injective⟩ +#align BddOrd.concrete_category BddOrd.concreteCategory + +instance hasForgetToPartOrd : HasForget₂ BddOrd PartOrd + where forget₂ := + { obj := fun X => X.toPartOrd + map := fun X Y => BoundedOrderHom.toOrderHom } +#align BddOrd.has_forget_to_PartOrd BddOrd.hasForgetToPartOrd + +instance hasForgetToBipointed : HasForget₂ BddOrd Bipointed + where + forget₂ := + { obj := fun X => ⟨X, ⊥, ⊤⟩ + map := fun X Y f => ⟨f, map_bot f, map_top f⟩ } + forget_comp := rfl +#align BddOrd.has_forget_to_Bipointed BddOrd.hasForgetToBipointed + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : BddOrd ⥤ BddOrd where + obj X := of Xᵒᵈ + map X Y := BoundedOrderHom.dual +#align BddOrd.dual BddOrd.dual + +/-- Constructs an equivalence between bounded orders from an order isomorphism between them. -/ +@[simps] +def Iso.mk {α β : BddOrd.{u}} (e : α ≃o β) : α ≅ β + where + Hom := e + inv := e.symm + hom_inv_id' := by + ext + exact e.symm_apply_apply _ + inv_hom_id' := by + ext + exact e.apply_symm_apply _ +#align BddOrd.iso.mk BddOrd.Iso.mk + +/-- The equivalence between `BddOrd` and itself induced by `order_dual` both ways. -/ +@[simps Functor inverse] +def dualEquiv : BddOrd ≌ BddOrd := + 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 BddOrd.dual_equiv BddOrd.dualEquiv + +end BddOrd + +theorem bddOrd_dual_comp_forget_to_partOrd : + BddOrd.dual ⋙ forget₂ BddOrd PartOrd = forget₂ BddOrd PartOrd ⋙ PartOrd.dual := + rfl +#align BddOrd_dual_comp_forget_to_PartOrd bddOrd_dual_comp_forget_to_partOrd + +theorem bddOrd_dual_comp_forget_to_bipointed : + BddOrd.dual ⋙ forget₂ BddOrd Bipointed = forget₂ BddOrd Bipointed ⋙ Bipointed.swap := + rfl +#align BddOrd_dual_comp_forget_to_Bipointed bddOrd_dual_comp_forget_to_bipointed + diff --git a/Mathbin/Order/Category/BoolAlg.lean b/Mathbin/Order/Category/BoolAlg.lean index 06c7031722..056e039567 100644 --- a/Mathbin/Order/Category/BoolAlg.lean +++ b/Mathbin/Order/Category/BoolAlg.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies ! This file was ported from Lean 3 source module order.category.BoolAlg -! leanprover-community/mathlib commit 51cbe88849c5bcf9eaf8e38f2bbdf1a44bbabe0c +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -49,25 +49,25 @@ theorem coe_of (α : Type _) [BooleanAlgebra α] : ↥(of α) = α := instance : Inhabited BoolAlg := ⟨of PUnit⟩ -/-- Turn a `BoolAlg` into a `BoundedDistribLattice` by forgetting its complement operation. -/ -def toBoundedDistribLattice (X : BoolAlg) : BoundedDistribLattice := - BoundedDistribLattice.of X -#align BoolAlg.to_BoundedDistribLattice BoolAlg.toBoundedDistribLattice +/-- Turn a `BoolAlg` into a `BddDistLat` by forgetting its complement operation. -/ +def toBddDistLat (X : BoolAlg) : BddDistLat := + BddDistLat.of X +#align BoolAlg.to_BddDistLat BoolAlg.toBddDistLat @[simp] -theorem coe_toBoundedDistribLattice (X : BoolAlg) : ↥X.toBoundedDistribLattice = ↥X := +theorem coe_toBddDistLat (X : BoolAlg) : ↥X.toBddDistLat = ↥X := rfl -#align BoolAlg.coe_to_BoundedDistribLattice BoolAlg.coe_toBoundedDistribLattice +#align BoolAlg.coe_to_BddDistLat BoolAlg.coe_toBddDistLat instance : LargeCategory.{u} BoolAlg := - InducedCategory.category toBoundedDistribLattice + InducedCategory.category toBddDistLat instance : ConcreteCategory BoolAlg := - InducedCategory.concreteCategory toBoundedDistribLattice + InducedCategory.concreteCategory toBddDistLat -instance hasForgetToBoundedDistribLattice : HasForget₂ BoolAlg BoundedDistribLattice := - InducedCategory.hasForget₂ toBoundedDistribLattice -#align BoolAlg.has_forget_to_BoundedDistribLattice BoolAlg.hasForgetToBoundedDistribLattice +instance hasForgetToBddDistLat : HasForget₂ BoolAlg BddDistLat := + InducedCategory.hasForget₂ toBddDistLat +#align BoolAlg.has_forget_to_BddDistLat BoolAlg.hasForgetToBddDistLat section @@ -113,9 +113,8 @@ def dualEquiv : BoolAlg ≌ BoolAlg := end BoolAlg -theorem boolAlg_dual_comp_forget_to_boundedDistribLattice : - BoolAlg.dual ⋙ forget₂ BoolAlg BoundedDistribLattice = - forget₂ BoolAlg BoundedDistribLattice ⋙ BoundedDistribLattice.dual := +theorem boolAlg_dual_comp_forget_to_bddDistLat : + BoolAlg.dual ⋙ forget₂ BoolAlg BddDistLat = forget₂ BoolAlg BddDistLat ⋙ BddDistLat.dual := rfl -#align BoolAlg_dual_comp_forget_to_BoundedDistribLattice boolAlg_dual_comp_forget_to_boundedDistribLattice +#align BoolAlg_dual_comp_forget_to_BddDistLat boolAlg_dual_comp_forget_to_bddDistLat diff --git a/Mathbin/Order/Category/BoundedDistribLattice.lean b/Mathbin/Order/Category/BoundedDistribLattice.lean deleted file mode 100644 index 465e76de86..0000000000 --- a/Mathbin/Order/Category/BoundedDistribLattice.lean +++ /dev/null @@ -1,127 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies - -! This file was ported from Lean 3 source module order.category.BoundedDistribLattice -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.Order.Category.BoundedLattice -import Mathbin.Order.Category.DistribLattice - -/-! -# The category of bounded distributive lattices - -This defines `BoundedDistribLattice`, the category of bounded distributive lattices. - -Note that this category is sometimes called [`DistLat`](https://ncatlab.org/nlab/show/DistLat) when -being a lattice is understood to entail having a bottom and a top element. --/ - - -universe u - -open CategoryTheory - -/-- The category of bounded distributive lattices with bounded lattice morphisms. -/ -structure BoundedDistribLattice where - toDistribLattice : DistribLatticeCat - [isBoundedOrder : BoundedOrder to_DistribLattice] -#align BoundedDistribLattice BoundedDistribLattice - -namespace BoundedDistribLattice - -instance : CoeSort BoundedDistribLattice (Type _) := - ⟨fun X => X.toDistribLattice⟩ - -instance (X : BoundedDistribLattice) : DistribLattice X := - X.toDistribLattice.str - -attribute [instance] BoundedDistribLattice.isBoundedOrder - -/-- Construct a bundled `BoundedDistribLattice` from a `bounded_order` `distrib_lattice`. -/ -def of (α : Type _) [DistribLattice α] [BoundedOrder α] : BoundedDistribLattice := - ⟨⟨α⟩⟩ -#align BoundedDistribLattice.of BoundedDistribLattice.of - -@[simp] -theorem coe_of (α : Type _) [DistribLattice α] [BoundedOrder α] : ↥(of α) = α := - rfl -#align BoundedDistribLattice.coe_of BoundedDistribLattice.coe_of - -instance : Inhabited BoundedDistribLattice := - ⟨of PUnit⟩ - -/-- Turn a `BoundedDistribLattice` into a `BoundedLattice` by forgetting it is distributive. -/ -def toBoundedLattice (X : BoundedDistribLattice) : BoundedLattice := - BoundedLattice.of X -#align BoundedDistribLattice.to_BoundedLattice BoundedDistribLattice.toBoundedLattice - -@[simp] -theorem coe_toBoundedLattice (X : BoundedDistribLattice) : ↥X.toBoundedLattice = ↥X := - rfl -#align BoundedDistribLattice.coe_to_BoundedLattice BoundedDistribLattice.coe_toBoundedLattice - -instance : LargeCategory.{u} BoundedDistribLattice := - InducedCategory.category toBoundedLattice - -instance : ConcreteCategory BoundedDistribLattice := - InducedCategory.concreteCategory toBoundedLattice - -instance hasForgetToDistribLattice : HasForget₂ BoundedDistribLattice DistribLatticeCat - where forget₂ := - { obj := fun X => ⟨X⟩ - map := fun X Y => BoundedLatticeHom.toLatticeHom } -#align BoundedDistribLattice.has_forget_to_DistribLattice BoundedDistribLattice.hasForgetToDistribLattice - -instance hasForgetToBoundedLattice : HasForget₂ BoundedDistribLattice BoundedLattice := - InducedCategory.hasForget₂ toBoundedLattice -#align BoundedDistribLattice.has_forget_to_BoundedLattice BoundedDistribLattice.hasForgetToBoundedLattice - -theorem forget_boundedLattice_latticeCat_eq_forget_distribLatticeCat_latticeCat : - forget₂ BoundedDistribLattice BoundedLattice ⋙ forget₂ BoundedLattice LatticeCat = - forget₂ BoundedDistribLattice DistribLatticeCat ⋙ forget₂ DistribLatticeCat LatticeCat := - rfl -#align BoundedDistribLattice.forget_BoundedLattice_Lattice_eq_forget_DistribLattice_Lattice BoundedDistribLattice.forget_boundedLattice_latticeCat_eq_forget_distribLatticeCat_latticeCat - -/-- Constructs an equivalence between bounded distributive lattices from an order isomorphism -between them. -/ -@[simps] -def Iso.mk {α β : BoundedDistribLattice.{u}} (e : α ≃o β) : α ≅ β - where - Hom := (e : BoundedLatticeHom α β) - inv := (e.symm : BoundedLatticeHom β α) - hom_inv_id' := by - ext - exact e.symm_apply_apply _ - inv_hom_id' := by - ext - exact e.apply_symm_apply _ -#align BoundedDistribLattice.iso.mk BoundedDistribLattice.Iso.mk - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : BoundedDistribLattice ⥤ BoundedDistribLattice - where - obj X := of Xᵒᵈ - map X Y := BoundedLatticeHom.dual -#align BoundedDistribLattice.dual BoundedDistribLattice.dual - -/-- The equivalence between `BoundedDistribLattice` and itself induced by `order_dual` both ways. -/ -@[simps Functor inverse] -def dualEquiv : BoundedDistribLattice ≌ BoundedDistribLattice := - 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 BoundedDistribLattice.dual_equiv BoundedDistribLattice.dualEquiv - -end BoundedDistribLattice - -theorem boundedDistribLattice_dual_comp_forget_to_distribLatticeCat : - BoundedDistribLattice.dual ⋙ forget₂ BoundedDistribLattice DistribLatticeCat = - forget₂ BoundedDistribLattice DistribLatticeCat ⋙ DistribLatticeCat.dual := - rfl -#align BoundedDistribLattice_dual_comp_forget_to_DistribLattice boundedDistribLattice_dual_comp_forget_to_distribLatticeCat - diff --git a/Mathbin/Order/Category/BoundedLattice.lean b/Mathbin/Order/Category/BoundedLattice.lean deleted file mode 100644 index 395d0651dd..0000000000 --- a/Mathbin/Order/Category/BoundedLattice.lean +++ /dev/null @@ -1,194 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies - -! This file was ported from Lean 3 source module order.category.BoundedLattice -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.Order.Category.BoundedOrder -import Mathbin.Order.Category.Lattice -import Mathbin.Order.Category.Semilattice - -/-! -# The category of bounded lattices - -This file defines `BoundedLattice`, the category of bounded lattices. - -In literature, this is sometimes called `Lat`, the category of lattices, because being a lattice is -understood to entail having a bottom and a top element. --/ - - -universe u - -open CategoryTheory - -/-- The category of bounded lattices with bounded lattice morphisms. -/ -structure BoundedLattice where - toLattice : LatticeCat - [isBoundedOrder : BoundedOrder to_Lattice] -#align BoundedLattice BoundedLattice - -namespace BoundedLattice - -instance : CoeSort BoundedLattice (Type _) := - ⟨fun X => X.toLattice⟩ - -instance (X : BoundedLattice) : Lattice X := - X.toLattice.str - -attribute [instance] BoundedLattice.isBoundedOrder - -/-- Construct a bundled `BoundedLattice` from `lattice` + `bounded_order`. -/ -def of (α : Type _) [Lattice α] [BoundedOrder α] : BoundedLattice := - ⟨⟨α⟩⟩ -#align BoundedLattice.of BoundedLattice.of - -@[simp] -theorem coe_of (α : Type _) [Lattice α] [BoundedOrder α] : ↥(of α) = α := - rfl -#align BoundedLattice.coe_of BoundedLattice.coe_of - -instance : Inhabited BoundedLattice := - ⟨of PUnit⟩ - -instance : LargeCategory.{u} BoundedLattice - where - Hom X Y := BoundedLatticeHom X Y - id X := BoundedLatticeHom.id X - comp X Y Z f g := g.comp f - id_comp' X Y := BoundedLatticeHom.comp_id - comp_id' X Y := BoundedLatticeHom.id_comp - assoc' W X Y Z _ _ _ := BoundedLatticeHom.comp_assoc _ _ _ - -instance : ConcreteCategory BoundedLattice - where - forget := ⟨coeSort, fun X Y => coeFn, fun X => rfl, fun X Y Z f g => rfl⟩ - forget_faithful := ⟨fun X Y => by convert FunLike.coe_injective⟩ - -instance hasForgetToBoundedOrder : HasForget₂ BoundedLattice BoundedOrderCat - where forget₂ := - { obj := fun X => BoundedOrderCat.of X - map := fun X Y => BoundedLatticeHom.toBoundedOrderHom } -#align BoundedLattice.has_forget_to_BoundedOrder BoundedLattice.hasForgetToBoundedOrder - -instance hasForgetToLattice : HasForget₂ BoundedLattice LatticeCat - where forget₂ := - { obj := fun X => ⟨X⟩ - map := fun X Y => BoundedLatticeHom.toLatticeHom } -#align BoundedLattice.has_forget_to_Lattice BoundedLattice.hasForgetToLattice - -instance hasForgetToSemilatticeSup : HasForget₂ BoundedLattice SemilatticeSupCat - where forget₂ := - { obj := fun X => ⟨X⟩ - map := fun X Y => BoundedLatticeHom.toSupBotHom } -#align BoundedLattice.has_forget_to_SemilatticeSup BoundedLattice.hasForgetToSemilatticeSup - -instance hasForgetToSemilatticeInf : HasForget₂ BoundedLattice SemilatticeInfCat - where forget₂ := - { obj := fun X => ⟨X⟩ - map := fun X Y => BoundedLatticeHom.toInfTopHom } -#align BoundedLattice.has_forget_to_SemilatticeInf BoundedLattice.hasForgetToSemilatticeInf - -@[simp] -theorem coe_forget_to_boundedOrderCat (X : BoundedLattice) : - ↥((forget₂ BoundedLattice BoundedOrderCat).obj X) = ↥X := - rfl -#align BoundedLattice.coe_forget_to_BoundedOrder BoundedLattice.coe_forget_to_boundedOrderCat - -@[simp] -theorem coe_forget_to_latticeCat (X : BoundedLattice) : - ↥((forget₂ BoundedLattice LatticeCat).obj X) = ↥X := - rfl -#align BoundedLattice.coe_forget_to_Lattice BoundedLattice.coe_forget_to_latticeCat - -@[simp] -theorem coe_forget_to_semilatticeSupCat (X : BoundedLattice) : - ↥((forget₂ BoundedLattice SemilatticeSupCat).obj X) = ↥X := - rfl -#align BoundedLattice.coe_forget_to_SemilatticeSup BoundedLattice.coe_forget_to_semilatticeSupCat - -@[simp] -theorem coe_forget_to_semilatticeInfCat (X : BoundedLattice) : - ↥((forget₂ BoundedLattice SemilatticeInfCat).obj X) = ↥X := - rfl -#align BoundedLattice.coe_forget_to_SemilatticeInf BoundedLattice.coe_forget_to_semilatticeInfCat - -theorem forget_latticeCat_partialOrderCat_eq_forget_boundedOrderCat_partialOrderCat : - forget₂ BoundedLattice LatticeCat ⋙ forget₂ LatticeCat PartialOrderCat = - forget₂ BoundedLattice BoundedOrderCat ⋙ forget₂ BoundedOrderCat PartialOrderCat := - rfl -#align BoundedLattice.forget_Lattice_PartialOrder_eq_forget_BoundedOrder_PartialOrder BoundedLattice.forget_latticeCat_partialOrderCat_eq_forget_boundedOrderCat_partialOrderCat - -theorem forget_semilatticeSupCat_partialOrderCat_eq_forget_boundedOrderCat_partialOrderCat : - forget₂ BoundedLattice SemilatticeSupCat ⋙ forget₂ SemilatticeSupCat PartialOrderCat = - forget₂ BoundedLattice BoundedOrderCat ⋙ forget₂ BoundedOrderCat PartialOrderCat := - rfl -#align BoundedLattice.forget_SemilatticeSup_PartialOrder_eq_forget_BoundedOrder_PartialOrder BoundedLattice.forget_semilatticeSupCat_partialOrderCat_eq_forget_boundedOrderCat_partialOrderCat - -theorem forget_semilatticeInfCat_partialOrderCat_eq_forget_boundedOrderCat_partialOrderCat : - forget₂ BoundedLattice SemilatticeInfCat ⋙ forget₂ SemilatticeInfCat PartialOrderCat = - forget₂ BoundedLattice BoundedOrderCat ⋙ forget₂ BoundedOrderCat PartialOrderCat := - rfl -#align BoundedLattice.forget_SemilatticeInf_PartialOrder_eq_forget_BoundedOrder_PartialOrder BoundedLattice.forget_semilatticeInfCat_partialOrderCat_eq_forget_boundedOrderCat_partialOrderCat - -/-- Constructs an equivalence between bounded lattices from an order isomorphism -between them. -/ -@[simps] -def Iso.mk {α β : BoundedLattice.{u}} (e : α ≃o β) : α ≅ β - where - Hom := e - inv := e.symm - hom_inv_id' := by - ext - exact e.symm_apply_apply _ - inv_hom_id' := by - ext - exact e.apply_symm_apply _ -#align BoundedLattice.iso.mk BoundedLattice.Iso.mk - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : BoundedLattice ⥤ BoundedLattice - where - obj X := of Xᵒᵈ - map X Y := BoundedLatticeHom.dual -#align BoundedLattice.dual BoundedLattice.dual - -/-- The equivalence between `BoundedLattice` and itself induced by `order_dual` both ways. -/ -@[simps Functor inverse] -def dualEquiv : BoundedLattice ≌ BoundedLattice := - 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 BoundedLattice.dual_equiv BoundedLattice.dualEquiv - -end BoundedLattice - -theorem boundedLattice_dual_comp_forget_to_boundedOrderCat : - BoundedLattice.dual ⋙ forget₂ BoundedLattice BoundedOrderCat = - forget₂ BoundedLattice BoundedOrderCat ⋙ BoundedOrderCat.dual := - rfl -#align BoundedLattice_dual_comp_forget_to_BoundedOrder boundedLattice_dual_comp_forget_to_boundedOrderCat - -theorem boundedLattice_dual_comp_forget_to_latticeCat : - BoundedLattice.dual ⋙ forget₂ BoundedLattice LatticeCat = - forget₂ BoundedLattice LatticeCat ⋙ LatticeCat.dual := - rfl -#align BoundedLattice_dual_comp_forget_to_Lattice boundedLattice_dual_comp_forget_to_latticeCat - -theorem boundedLattice_dual_comp_forget_to_semilatticeSupCat : - BoundedLattice.dual ⋙ forget₂ BoundedLattice SemilatticeSupCat = - forget₂ BoundedLattice SemilatticeInfCat ⋙ SemilatticeInfCat.dual := - rfl -#align BoundedLattice_dual_comp_forget_to_SemilatticeSup boundedLattice_dual_comp_forget_to_semilatticeSupCat - -theorem boundedLattice_dual_comp_forget_to_semilatticeInfCat : - BoundedLattice.dual ⋙ forget₂ BoundedLattice SemilatticeInfCat = - forget₂ BoundedLattice SemilatticeSupCat ⋙ SemilatticeSupCat.dual := - rfl -#align BoundedLattice_dual_comp_forget_to_SemilatticeInf boundedLattice_dual_comp_forget_to_semilatticeInfCat - diff --git a/Mathbin/Order/Category/BoundedOrder.lean b/Mathbin/Order/Category/BoundedOrder.lean deleted file mode 100644 index b0ff826d9a..0000000000 --- a/Mathbin/Order/Category/BoundedOrder.lean +++ /dev/null @@ -1,128 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies - -! This file was ported from Lean 3 source module order.category.BoundedOrder -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.CategoryTheory.Category.Bipointed -import Mathbin.Order.Category.PartialOrder -import Mathbin.Order.Hom.Bounded - -/-! -# The category of bounded orders - -This defines `BoundedOrder`, the category of bounded orders. --/ - - -universe u v - -open CategoryTheory - -/-- The category of bounded orders with monotone functions. -/ -structure BoundedOrderCat where - toPartialOrder : PartialOrderCat - [isBoundedOrder : BoundedOrder to_PartialOrder] -#align BoundedOrder BoundedOrderCat - -namespace BoundedOrderCat - -instance : CoeSort BoundedOrderCat (Type _) := - InducedCategory.hasCoeToSort toPartialOrder - -instance (X : BoundedOrderCat) : PartialOrder X := - X.toPartialOrder.str - -attribute [instance] BoundedOrderCat.isBoundedOrder - -/-- Construct a bundled `BoundedOrder` from a `fintype` `partial_order`. -/ -def of (α : Type _) [PartialOrder α] [BoundedOrder α] : BoundedOrderCat := - ⟨⟨α⟩⟩ -#align BoundedOrder.of BoundedOrderCat.of - -@[simp] -theorem coe_of (α : Type _) [PartialOrder α] [BoundedOrder α] : ↥(of α) = α := - rfl -#align BoundedOrder.coe_of BoundedOrderCat.coe_of - -instance : Inhabited BoundedOrderCat := - ⟨of PUnit⟩ - -instance largeCategory : LargeCategory.{u} BoundedOrderCat - where - Hom X Y := BoundedOrderHom X Y - id X := BoundedOrderHom.id X - comp X Y Z f g := g.comp f - id_comp' X Y := BoundedOrderHom.comp_id - comp_id' X Y := BoundedOrderHom.id_comp - assoc' W X Y Z _ _ _ := BoundedOrderHom.comp_assoc _ _ _ -#align BoundedOrder.large_category BoundedOrderCat.largeCategory - -instance concreteCategory : ConcreteCategory BoundedOrderCat - where - forget := ⟨coeSort, fun X Y => coeFn, fun X => rfl, fun X Y Z f g => rfl⟩ - forget_faithful := ⟨fun X Y => by convert FunLike.coe_injective⟩ -#align BoundedOrder.concrete_category BoundedOrderCat.concreteCategory - -instance hasForgetToPartialOrder : HasForget₂ BoundedOrderCat PartialOrderCat - where forget₂ := - { obj := fun X => X.toPartialOrder - map := fun X Y => BoundedOrderHom.toOrderHom } -#align BoundedOrder.has_forget_to_PartialOrder BoundedOrderCat.hasForgetToPartialOrder - -instance hasForgetToBipointed : HasForget₂ BoundedOrderCat Bipointed - where - forget₂ := - { obj := fun X => ⟨X, ⊥, ⊤⟩ - map := fun X Y f => ⟨f, map_bot f, map_top f⟩ } - forget_comp := rfl -#align BoundedOrder.has_forget_to_Bipointed BoundedOrderCat.hasForgetToBipointed - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : BoundedOrderCat ⥤ BoundedOrderCat - where - obj X := of Xᵒᵈ - map X Y := BoundedOrderHom.dual -#align BoundedOrder.dual BoundedOrderCat.dual - -/-- Constructs an equivalence between bounded orders from an order isomorphism between them. -/ -@[simps] -def Iso.mk {α β : BoundedOrderCat.{u}} (e : α ≃o β) : α ≅ β - where - Hom := e - inv := e.symm - hom_inv_id' := by - ext - exact e.symm_apply_apply _ - inv_hom_id' := by - ext - exact e.apply_symm_apply _ -#align BoundedOrder.iso.mk BoundedOrderCat.Iso.mk - -/-- The equivalence between `BoundedOrder` and itself induced by `order_dual` both ways. -/ -@[simps Functor inverse] -def dualEquiv : BoundedOrderCat ≌ BoundedOrderCat := - 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 BoundedOrder.dual_equiv BoundedOrderCat.dualEquiv - -end BoundedOrderCat - -theorem boundedOrderCat_dual_comp_forget_to_partialOrderCat : - BoundedOrderCat.dual ⋙ forget₂ BoundedOrderCat PartialOrderCat = - forget₂ BoundedOrderCat PartialOrderCat ⋙ PartialOrderCat.dual := - rfl -#align BoundedOrder_dual_comp_forget_to_PartialOrder boundedOrderCat_dual_comp_forget_to_partialOrderCat - -theorem boundedOrderCat_dual_comp_forget_to_bipointed : - BoundedOrderCat.dual ⋙ forget₂ BoundedOrderCat Bipointed = - forget₂ BoundedOrderCat Bipointed ⋙ Bipointed.swap := - rfl -#align BoundedOrder_dual_comp_forget_to_Bipointed boundedOrderCat_dual_comp_forget_to_bipointed - diff --git a/Mathbin/Order/Category/CompleteLat.lean b/Mathbin/Order/Category/CompleteLat.lean new file mode 100644 index 0000000000..1adcc8bc51 --- /dev/null +++ b/Mathbin/Order/Category/CompleteLat.lean @@ -0,0 +1,107 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies + +! This file was ported from Lean 3 source module order.category.CompleteLat +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Order.Category.BddLat +import Mathbin.Order.Hom.CompleteLattice + +/-! +# The category of complete lattices + +This file defines `CompleteLat`, the category of complete lattices. +-/ + + +universe u + +open CategoryTheory + +/-- The category of complete lattices. -/ +def CompleteLat := + Bundled CompleteLattice +#align CompleteLat CompleteLat + +namespace CompleteLat + +instance : CoeSort CompleteLat (Type _) := + Bundled.hasCoeToSort + +instance (X : CompleteLat) : CompleteLattice X := + X.str + +/-- Construct a bundled `CompleteLat` from a `complete_lattice`. -/ +def of (α : Type _) [CompleteLattice α] : CompleteLat := + Bundled.of α +#align CompleteLat.of CompleteLat.of + +@[simp] +theorem coe_of (α : Type _) [CompleteLattice α] : ↥(of α) = α := + rfl +#align CompleteLat.coe_of CompleteLat.coe_of + +instance : Inhabited CompleteLat := + ⟨of PUnit⟩ + +instance : BundledHom @CompleteLatticeHom + where + toFun _ _ _ _ := coeFn + id := @CompleteLatticeHom.id + comp := @CompleteLatticeHom.comp + hom_ext X Y _ _ := FunLike.coe_injective + +instance : LargeCategory.{u} CompleteLat := + BundledHom.category CompleteLatticeHom + +instance : ConcreteCategory CompleteLat := + BundledHom.concreteCategory CompleteLatticeHom + +instance hasForgetToBddLat : HasForget₂ CompleteLat BddLat + where + forget₂ := + { obj := fun X => BddLat.of X + map := fun X Y => CompleteLatticeHom.toBoundedLatticeHom } + forget_comp := rfl +#align CompleteLat.has_forget_to_BddLat CompleteLat.hasForgetToBddLat + +/-- Constructs an isomorphism of complete lattices from an order isomorphism between them. -/ +@[simps] +def Iso.mk {α β : CompleteLat.{u}} (e : α ≃o β) : α ≅ β + where + Hom := e + inv := e.symm + hom_inv_id' := by + ext + exact e.symm_apply_apply _ + inv_hom_id' := by + ext + exact e.apply_symm_apply _ +#align CompleteLat.iso.mk CompleteLat.Iso.mk + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : CompleteLat ⥤ CompleteLat where + obj X := of Xᵒᵈ + map X Y := CompleteLatticeHom.dual +#align CompleteLat.dual CompleteLat.dual + +/-- The equivalence between `CompleteLat` and itself induced by `order_dual` both ways. -/ +@[simps Functor inverse] +def dualEquiv : CompleteLat ≌ CompleteLat := + 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 CompleteLat.dual_equiv CompleteLat.dualEquiv + +end CompleteLat + +theorem completeLat_dual_comp_forget_to_bddLat : + CompleteLat.dual ⋙ forget₂ CompleteLat BddLat = forget₂ CompleteLat BddLat ⋙ BddLat.dual := + rfl +#align CompleteLat_dual_comp_forget_to_BddLat completeLat_dual_comp_forget_to_bddLat + diff --git a/Mathbin/Order/Category/CompleteLattice.lean b/Mathbin/Order/Category/CompleteLattice.lean deleted file mode 100644 index 5125e53c42..0000000000 --- a/Mathbin/Order/Category/CompleteLattice.lean +++ /dev/null @@ -1,109 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies - -! This file was ported from Lean 3 source module order.category.CompleteLattice -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.Order.Category.BoundedLattice -import Mathbin.Order.Hom.CompleteLattice - -/-! -# The category of complete lattices - -This file defines `CompleteLattice`, the category of complete lattices. --/ - - -universe u - -open CategoryTheory - -/-- The category of complete lattices. -/ -def CompleteLatticeCat := - Bundled CompleteLattice -#align CompleteLattice CompleteLatticeCat - -namespace CompleteLatticeCat - -instance : CoeSort CompleteLatticeCat (Type _) := - Bundled.hasCoeToSort - -instance (X : CompleteLatticeCat) : CompleteLattice X := - X.str - -/-- Construct a bundled `CompleteLattice` from a `complete_lattice`. -/ -def of (α : Type _) [CompleteLattice α] : CompleteLatticeCat := - Bundled.of α -#align CompleteLattice.of CompleteLatticeCat.of - -@[simp] -theorem coe_of (α : Type _) [CompleteLattice α] : ↥(of α) = α := - rfl -#align CompleteLattice.coe_of CompleteLatticeCat.coe_of - -instance : Inhabited CompleteLatticeCat := - ⟨of PUnit⟩ - -instance : BundledHom @CompleteLatticeHom - where - toFun _ _ _ _ := coeFn - id := @CompleteLatticeHom.id - comp := @CompleteLatticeHom.comp - hom_ext X Y _ _ := FunLike.coe_injective - -instance : LargeCategory.{u} CompleteLatticeCat := - BundledHom.category CompleteLatticeHom - -instance : ConcreteCategory CompleteLatticeCat := - BundledHom.concreteCategory CompleteLatticeHom - -instance hasForgetToBoundedLattice : HasForget₂ CompleteLatticeCat BoundedLattice - where - forget₂ := - { obj := fun X => BoundedLattice.of X - map := fun X Y => CompleteLatticeHom.toBoundedLatticeHom } - forget_comp := rfl -#align CompleteLattice.has_forget_to_BoundedLattice CompleteLatticeCat.hasForgetToBoundedLattice - -/-- Constructs an isomorphism of complete lattices from an order isomorphism between them. -/ -@[simps] -def Iso.mk {α β : CompleteLatticeCat.{u}} (e : α ≃o β) : α ≅ β - where - Hom := e - inv := e.symm - hom_inv_id' := by - ext - exact e.symm_apply_apply _ - inv_hom_id' := by - ext - exact e.apply_symm_apply _ -#align CompleteLattice.iso.mk CompleteLatticeCat.Iso.mk - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : CompleteLatticeCat ⥤ CompleteLatticeCat - where - obj X := of Xᵒᵈ - map X Y := CompleteLatticeHom.dual -#align CompleteLattice.dual CompleteLatticeCat.dual - -/-- The equivalence between `CompleteLattice` and itself induced by `order_dual` both ways. -/ -@[simps Functor inverse] -def dualEquiv : CompleteLatticeCat ≌ CompleteLatticeCat := - 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 CompleteLattice.dual_equiv CompleteLatticeCat.dualEquiv - -end CompleteLatticeCat - -theorem completeLatticeCat_dual_comp_forget_to_boundedLattice : - CompleteLatticeCat.dual ⋙ forget₂ CompleteLatticeCat BoundedLattice = - forget₂ CompleteLatticeCat BoundedLattice ⋙ BoundedLattice.dual := - rfl -#align CompleteLattice_dual_comp_forget_to_BoundedLattice completeLatticeCat_dual_comp_forget_to_boundedLattice - diff --git a/Mathbin/Order/Category/DistLat.lean b/Mathbin/Order/Category/DistLat.lean new file mode 100644 index 0000000000..a3d43cd980 --- /dev/null +++ b/Mathbin/Order/Category/DistLat.lean @@ -0,0 +1,99 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies + +! This file was ported from Lean 3 source module order.category.DistLat +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Order.Category.Lat + +/-! +# The category of distributive lattices + +This file defines `DistLat`, the category of distributive lattices. + +Note that [`DistLat`](https://ncatlab.org/nlab/show/DistLat) in the literature doesn't always +correspond to `DistLat` as we don't require bottom or top elements. Instead, this `DistLat` +corresponds to `BddDistLat`. +-/ + + +universe u + +open CategoryTheory + +/-- The category of distributive lattices. -/ +def DistLat := + Bundled DistribLattice +#align DistLat DistLat + +namespace DistLat + +instance : CoeSort DistLat (Type _) := + Bundled.hasCoeToSort + +instance (X : DistLat) : DistribLattice X := + X.str + +/-- Construct a bundled `DistLat` from a `distrib_lattice` underlying type and typeclass. -/ +def of (α : Type _) [DistribLattice α] : DistLat := + Bundled.of α +#align DistLat.of DistLat.of + +@[simp] +theorem coe_of (α : Type _) [DistribLattice α] : ↥(of α) = α := + rfl +#align DistLat.coe_of DistLat.coe_of + +instance : Inhabited DistLat := + ⟨of PUnit⟩ + +instance : BundledHom.ParentProjection @DistribLattice.toLattice := + ⟨⟩ + +deriving instance LargeCategory, ConcreteCategory for DistLat + +instance hasForgetToLat : HasForget₂ DistLat Lat := + BundledHom.forget₂ _ _ +#align DistLat.has_forget_to_Lat DistLat.hasForgetToLat + +/-- Constructs an equivalence between distributive lattices from an order isomorphism between them. +-/ +@[simps] +def Iso.mk {α β : DistLat.{u}} (e : α ≃o β) : α ≅ β + where + Hom := e + inv := e.symm + hom_inv_id' := by + ext + exact e.symm_apply_apply _ + inv_hom_id' := by + ext + exact e.apply_symm_apply _ +#align DistLat.iso.mk DistLat.Iso.mk + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : DistLat ⥤ DistLat where + obj X := of Xᵒᵈ + map X Y := LatticeHom.dual +#align DistLat.dual DistLat.dual + +/-- The equivalence between `DistLat` and itself induced by `order_dual` both ways. -/ +@[simps Functor inverse] +def dualEquiv : DistLat ≌ DistLat := + 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_lat : + DistLat.dual ⋙ forget₂ DistLat Lat = forget₂ DistLat Lat ⋙ Lat.dual := + rfl +#align DistLat_dual_comp_forget_to_Lat distLat_dual_comp_forget_to_lat + diff --git a/Mathbin/Order/Category/DistribLattice.lean b/Mathbin/Order/Category/DistribLattice.lean deleted file mode 100644 index ae327cd396..0000000000 --- a/Mathbin/Order/Category/DistribLattice.lean +++ /dev/null @@ -1,101 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies - -! This file was ported from Lean 3 source module order.category.DistribLattice -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.Order.Category.Lattice - -/-! -# The category of distributive lattices - -This file defines `DistribLattice`, the category of distributive lattices. - -Note that [`DistLat`](https://ncatlab.org/nlab/show/DistLat) in the literature doesn't always -correspond to `DistribLattice` as we don't require bottom or top elements. Instead, this `DistLat` -corresponds to `BoundedDistribLattice`. --/ - - -universe u - -open CategoryTheory - -/-- The category of distributive lattices. -/ -def DistribLatticeCat := - Bundled DistribLattice -#align DistribLattice DistribLatticeCat - -namespace DistribLatticeCat - -instance : CoeSort DistribLatticeCat (Type _) := - Bundled.hasCoeToSort - -instance (X : DistribLatticeCat) : DistribLattice X := - X.str - -/-- Construct a bundled `DistribLattice` from a `distrib_lattice` underlying type and typeclass. -/ -def of (α : Type _) [DistribLattice α] : DistribLatticeCat := - Bundled.of α -#align DistribLattice.of DistribLatticeCat.of - -@[simp] -theorem coe_of (α : Type _) [DistribLattice α] : ↥(of α) = α := - rfl -#align DistribLattice.coe_of DistribLatticeCat.coe_of - -instance : Inhabited DistribLatticeCat := - ⟨of PUnit⟩ - -instance : BundledHom.ParentProjection @DistribLattice.toLattice := - ⟨⟩ - -deriving instance LargeCategory, ConcreteCategory for DistribLatticeCat - -instance hasForgetToLattice : HasForget₂ DistribLatticeCat LatticeCat := - BundledHom.forget₂ _ _ -#align DistribLattice.has_forget_to_Lattice DistribLatticeCat.hasForgetToLattice - -/-- Constructs an equivalence between distributive lattices from an order isomorphism between them. --/ -@[simps] -def Iso.mk {α β : DistribLatticeCat.{u}} (e : α ≃o β) : α ≅ β - where - Hom := e - inv := e.symm - hom_inv_id' := by - ext - exact e.symm_apply_apply _ - inv_hom_id' := by - ext - exact e.apply_symm_apply _ -#align DistribLattice.iso.mk DistribLatticeCat.Iso.mk - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : DistribLatticeCat ⥤ DistribLatticeCat - where - obj X := of Xᵒᵈ - map X Y := LatticeHom.dual -#align DistribLattice.dual DistribLatticeCat.dual - -/-- The equivalence between `DistribLattice` and itself induced by `order_dual` both ways. -/ -@[simps Functor inverse] -def dualEquiv : DistribLatticeCat ≌ DistribLatticeCat := - 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 DistribLattice.dual_equiv DistribLatticeCat.dualEquiv - -end DistribLatticeCat - -theorem distribLatticeCat_dual_comp_forget_to_latticeCat : - DistribLatticeCat.dual ⋙ forget₂ DistribLatticeCat LatticeCat = - forget₂ DistribLatticeCat LatticeCat ⋙ LatticeCat.dual := - rfl -#align DistribLattice_dual_comp_forget_to_Lattice distribLatticeCat_dual_comp_forget_to_latticeCat - diff --git a/Mathbin/Order/Category/FinBoolAlg.lean b/Mathbin/Order/Category/FinBoolAlg.lean index 94f6084d67..5e5a2c5950 100644 --- a/Mathbin/Order/Category/FinBoolAlg.lean +++ b/Mathbin/Order/Category/FinBoolAlg.lean @@ -4,13 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies ! This file was ported from Lean 3 source module order.category.FinBoolAlg -! leanprover-community/mathlib commit c3019c79074b0619edb4b27553a91b2e82242395 +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ import Mathbin.Data.Fintype.Powerset import Mathbin.Order.Category.BoolAlg -import Mathbin.Order.Category.FinPartialOrder +import Mathbin.Order.Category.FinPartOrd import Mathbin.Order.Hom.CompleteLattice /-! @@ -87,17 +87,17 @@ instance forget_to_boolAlg_faithful : Faithful (forget₂ FinBoolAlg BoolAlg) := #align FinBoolAlg.forget_to_BoolAlg_faithful FinBoolAlg.forget_to_boolAlg_faithful @[simps] -instance hasForgetToFinPartialOrder : HasForget₂ FinBoolAlg FinPartialOrder +instance hasForgetToFinPartOrd : HasForget₂ FinBoolAlg FinPartOrd where forget₂ := - { obj := fun X => FinPartialOrder.of X + { obj := fun X => FinPartOrd.of X map := fun X Y f => show OrderHom X Y from ↑(show BoundedLatticeHom X Y from f) } -#align FinBoolAlg.has_forget_to_FinPartialOrder FinBoolAlg.hasForgetToFinPartialOrder +#align FinBoolAlg.has_forget_to_FinPartOrd FinBoolAlg.hasForgetToFinPartOrd -instance forget_to_finPartialOrder_faithful : Faithful (forget₂ FinBoolAlg FinPartialOrder) := +instance forget_to_finPartOrd_faithful : Faithful (forget₂ FinBoolAlg FinPartOrd) := ⟨fun X Y f g h => haveI := congr_arg (coeFn : _ → X → Y) h FunLike.coe_injective this⟩ -#align FinBoolAlg.forget_to_FinPartialOrder_faithful FinBoolAlg.forget_to_finPartialOrder_faithful +#align FinBoolAlg.forget_to_FinPartOrd_faithful FinBoolAlg.forget_to_finPartOrd_faithful /-- Constructs an equivalence between finite Boolean algebras from an order isomorphism between them. -/ diff --git a/Mathbin/Order/Category/FinPartOrd.lean b/Mathbin/Order/Category/FinPartOrd.lean new file mode 100644 index 0000000000..a2dc0cf487 --- /dev/null +++ b/Mathbin/Order/Category/FinPartOrd.lean @@ -0,0 +1,119 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies + +! This file was ported from Lean 3 source module order.category.FinPartOrd +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Fintype +import Mathbin.Order.Category.PartOrd + +/-! +# The category of finite partial orders + +This defines `FinPartOrd`, the category of finite partial orders. + +Note: `FinPartOrd` is NOT a subcategory of `BddOrd` because its morphisms do not +preserve `⊥` and `⊤`. + +## TODO + +`FinPartOrd` is equivalent to a small category. +-/ + + +universe u v + +open CategoryTheory + +/-- The category of finite partial orders with monotone functions. -/ +structure FinPartOrd where + toPartOrd : PartOrd + [isFintype : Fintype to_PartOrd] +#align FinPartOrd FinPartOrd + +namespace FinPartOrd + +instance : CoeSort FinPartOrd (Type _) := + ⟨fun X => X.toPartOrd⟩ + +instance (X : FinPartOrd) : PartialOrder X := + X.toPartOrd.str + +attribute [instance] FinPartOrd.isFintype + +@[simp] +theorem coe_toPartOrd (X : FinPartOrd) : ↥X.toPartOrd = ↥X := + rfl +#align FinPartOrd.coe_to_PartOrd FinPartOrd.coe_toPartOrd + +/-- Construct a bundled `FinPartOrd` from `fintype` + `partial_order`. -/ +def of (α : Type _) [PartialOrder α] [Fintype α] : FinPartOrd := + ⟨⟨α⟩⟩ +#align FinPartOrd.of FinPartOrd.of + +@[simp] +theorem coe_of (α : Type _) [PartialOrder α] [Fintype α] : ↥(of α) = α := + rfl +#align FinPartOrd.coe_of FinPartOrd.coe_of + +instance : Inhabited FinPartOrd := + ⟨of PUnit⟩ + +instance largeCategory : LargeCategory FinPartOrd := + InducedCategory.category FinPartOrd.toPartOrd +#align FinPartOrd.large_category FinPartOrd.largeCategory + +instance concreteCategory : ConcreteCategory FinPartOrd := + InducedCategory.concreteCategory FinPartOrd.toPartOrd +#align FinPartOrd.concrete_category FinPartOrd.concreteCategory + +instance hasForgetToPartOrd : HasForget₂ FinPartOrd PartOrd := + InducedCategory.hasForget₂ FinPartOrd.toPartOrd +#align FinPartOrd.has_forget_to_PartOrd FinPartOrd.hasForgetToPartOrd + +instance hasForgetToFintype : HasForget₂ FinPartOrd FintypeCat + where forget₂ := + { obj := fun X => ⟨X⟩ + map := fun X Y => coeFn } +#align FinPartOrd.has_forget_to_Fintype FinPartOrd.hasForgetToFintype + +/-- Constructs an isomorphism of finite partial orders from an order isomorphism between them. -/ +@[simps] +def Iso.mk {α β : FinPartOrd.{u}} (e : α ≃o β) : α ≅ β + where + Hom := e + inv := e.symm + hom_inv_id' := by + ext + exact e.symm_apply_apply _ + inv_hom_id' := by + ext + exact e.apply_symm_apply _ +#align FinPartOrd.iso.mk FinPartOrd.Iso.mk + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : FinPartOrd ⥤ FinPartOrd where + obj X := of Xᵒᵈ + map X Y := OrderHom.dual +#align FinPartOrd.dual FinPartOrd.dual + +/-- The equivalence between `FinPartOrd` and itself induced by `order_dual` both ways. -/ +@[simps Functor inverse] +def dualEquiv : FinPartOrd ≌ FinPartOrd := + 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 FinPartOrd.dual_equiv FinPartOrd.dualEquiv + +end FinPartOrd + +theorem finPartOrd_dual_comp_forget_to_partOrd : + FinPartOrd.dual ⋙ forget₂ FinPartOrd PartOrd = forget₂ FinPartOrd PartOrd ⋙ PartOrd.dual := + rfl +#align FinPartOrd_dual_comp_forget_to_PartOrd finPartOrd_dual_comp_forget_to_partOrd + diff --git a/Mathbin/Order/Category/FinPartialOrder.lean b/Mathbin/Order/Category/FinPartialOrder.lean deleted file mode 100644 index f4d6b669fa..0000000000 --- a/Mathbin/Order/Category/FinPartialOrder.lean +++ /dev/null @@ -1,121 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies - -! This file was ported from Lean 3 source module order.category.FinPartialOrder -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.CategoryTheory.Fintype -import Mathbin.Order.Category.PartialOrder - -/-! -# The category of finite partial orders - -This defines `FinPartialOrder`, the category of finite partial orders. - -Note: `FinPartialOrder` is NOT a subcategory of `BoundedOrder` because its morphisms do not -preserve `⊥` and `⊤`. - -## TODO - -`FinPartialOrder` is equivalent to a small category. --/ - - -universe u v - -open CategoryTheory - -/-- The category of finite partial orders with monotone functions. -/ -structure FinPartialOrder where - toPartialOrder : PartialOrderCat - [isFintype : Fintype to_PartialOrder] -#align FinPartialOrder FinPartialOrder - -namespace FinPartialOrder - -instance : CoeSort FinPartialOrder (Type _) := - ⟨fun X => X.toPartialOrder⟩ - -instance (X : FinPartialOrder) : PartialOrder X := - X.toPartialOrder.str - -attribute [instance] FinPartialOrder.isFintype - -@[simp] -theorem coe_toPartialOrder (X : FinPartialOrder) : ↥X.toPartialOrder = ↥X := - rfl -#align FinPartialOrder.coe_to_PartialOrder FinPartialOrder.coe_toPartialOrder - -/-- Construct a bundled `FinPartialOrder` from `fintype` + `partial_order`. -/ -def of (α : Type _) [PartialOrder α] [Fintype α] : FinPartialOrder := - ⟨⟨α⟩⟩ -#align FinPartialOrder.of FinPartialOrder.of - -@[simp] -theorem coe_of (α : Type _) [PartialOrder α] [Fintype α] : ↥(of α) = α := - rfl -#align FinPartialOrder.coe_of FinPartialOrder.coe_of - -instance : Inhabited FinPartialOrder := - ⟨of PUnit⟩ - -instance largeCategory : LargeCategory FinPartialOrder := - InducedCategory.category FinPartialOrder.toPartialOrder -#align FinPartialOrder.large_category FinPartialOrder.largeCategory - -instance concreteCategory : ConcreteCategory FinPartialOrder := - InducedCategory.concreteCategory FinPartialOrder.toPartialOrder -#align FinPartialOrder.concrete_category FinPartialOrder.concreteCategory - -instance hasForgetToPartialOrder : HasForget₂ FinPartialOrder PartialOrderCat := - InducedCategory.hasForget₂ FinPartialOrder.toPartialOrder -#align FinPartialOrder.has_forget_to_PartialOrder FinPartialOrder.hasForgetToPartialOrder - -instance hasForgetToFintype : HasForget₂ FinPartialOrder FintypeCat - where forget₂ := - { obj := fun X => ⟨X⟩ - map := fun X Y => coeFn } -#align FinPartialOrder.has_forget_to_Fintype FinPartialOrder.hasForgetToFintype - -/-- Constructs an isomorphism of finite partial orders from an order isomorphism between them. -/ -@[simps] -def Iso.mk {α β : FinPartialOrder.{u}} (e : α ≃o β) : α ≅ β - where - Hom := e - inv := e.symm - hom_inv_id' := by - ext - exact e.symm_apply_apply _ - inv_hom_id' := by - ext - exact e.apply_symm_apply _ -#align FinPartialOrder.iso.mk FinPartialOrder.Iso.mk - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : FinPartialOrder ⥤ FinPartialOrder - where - obj X := of Xᵒᵈ - map X Y := OrderHom.dual -#align FinPartialOrder.dual FinPartialOrder.dual - -/-- The equivalence between `FinPartialOrder` and itself induced by `order_dual` both ways. -/ -@[simps Functor inverse] -def dualEquiv : FinPartialOrder ≌ FinPartialOrder := - 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 FinPartialOrder.dual_equiv FinPartialOrder.dualEquiv - -end FinPartialOrder - -theorem finPartialOrder_dual_comp_forget_to_partialOrderCat : - FinPartialOrder.dual ⋙ forget₂ FinPartialOrder PartialOrderCat = - forget₂ FinPartialOrder PartialOrderCat ⋙ PartialOrderCat.dual := - rfl -#align FinPartialOrder_dual_comp_forget_to_PartialOrder finPartialOrder_dual_comp_forget_to_partialOrderCat - diff --git a/Mathbin/Order/Category/Frame.lean b/Mathbin/Order/Category/Frm.lean similarity index 68% rename from Mathbin/Order/Category/Frame.lean rename to Mathbin/Order/Category/Frm.lean index 0d63a56044..86283a03fd 100644 --- a/Mathbin/Order/Category/Frame.lean +++ b/Mathbin/Order/Category/Frm.lean @@ -3,12 +3,12 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -! This file was ported from Lean 3 source module order.category.Frame -! leanprover-community/mathlib commit 829895f162a1f29d0133f4b3538f4cd1fb5bffd3 +! This file was ported from Lean 3 source module order.category.Frm +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.Order.Category.Lattice +import Mathbin.Order.Category.Lat import Mathbin.Order.Hom.CompleteLattice import Mathbin.Topology.Category.CompHaus.Basic import Mathbin.Topology.Sets.Opens @@ -16,7 +16,7 @@ import Mathbin.Topology.Sets.Opens /-! # The category of frames -This file defines `Frame`, the category of frames. +This file defines `Frm`, the category of frames. ## References @@ -29,54 +29,54 @@ universe u open CategoryTheory Opposite Order TopologicalSpace /-- The category of frames. -/ -def Frame := +def Frm := Bundled Frame -#align Frame Frame +#align Frm Frm -namespace Frame +namespace Frm -instance : CoeSort Frame (Type _) := +instance : CoeSort Frm (Type _) := Bundled.hasCoeToSort -instance (X : Frame) : Frame X := +instance (X : Frm) : Frame X := X.str -/-- Construct a bundled `Frame` from a `frame`. -/ -def of (α : Type _) [Frame α] : Frame := +/-- Construct a bundled `Frm` from a `frame`. -/ +def of (α : Type _) [Frame α] : Frm := Bundled.of α -#align Frame.of Frame.of +#align Frm.of Frm.of @[simp] theorem coe_of (α : Type _) [Frame α] : ↥(of α) = α := rfl -#align Frame.coe_of Frame.coe_of +#align Frm.coe_of Frm.coe_of -instance : Inhabited Frame := +instance : Inhabited Frm := ⟨of PUnit⟩ /-- An abbreviation of `frame_hom` that assumes `frame` instead of the weaker `complete_lattice`. Necessary for the category theory machinery. -/ abbrev Hom (α β : Type _) [Frame α] [Frame β] : Type _ := FrameHom α β -#align Frame.hom Frame.Hom +#align Frm.hom Frm.Hom instance bundledHom : BundledHom Hom := ⟨fun α β [Frame α] [Frame β] => (coeFn : FrameHom α β → α → β), fun α [Frame α] => FrameHom.id α, fun α β γ [Frame α] [Frame β] [Frame γ] => FrameHom.comp, fun α β [Frame α] [Frame β] => FunLike.coe_injective⟩ -#align Frame.bundled_hom Frame.bundledHom +#align Frm.bundled_hom Frm.bundledHom -deriving instance LargeCategory, ConcreteCategory for Frame +deriving instance LargeCategory, ConcreteCategory for Frm -instance hasForgetToLattice : HasForget₂ Frame LatticeCat +instance hasForgetToLat : HasForget₂ Frm Lat where forget₂ := { obj := fun X => ⟨X⟩ map := fun X Y => FrameHom.toLatticeHom } -#align Frame.has_forget_to_Lattice Frame.hasForgetToLattice +#align Frm.has_forget_to_Lat Frm.hasForgetToLat /-- Constructs an isomorphism of frames from an order isomorphism between them. -/ @[simps] -def Iso.mk {α β : Frame.{u}} (e : α ≃o β) : α ≅ β +def Iso.mk {α β : Frm.{u}} (e : α ≃o β) : α ≅ β where Hom := e inv := e.symm @@ -86,15 +86,15 @@ def Iso.mk {α β : Frame.{u}} (e : α ≃o β) : α ≅ β inv_hom_id' := by ext exact e.apply_symm_apply _ -#align Frame.iso.mk Frame.Iso.mk +#align Frm.iso.mk Frm.Iso.mk -end Frame +end Frm -/-- The forgetful functor from `Topᵒᵖ` to `Frame`. -/ +/-- The forgetful functor from `Topᵒᵖ` to `Frm`. -/ @[simps] -def topOpToFrame : TopCatᵒᵖ ⥤ Frame +def topOpToFrame : TopCatᵒᵖ ⥤ Frm where - obj X := Frame.of (Opens (unop X : TopCat)) + obj X := Frm.of (Opens (unop X : TopCat)) map X Y f := Opens.comap <| Quiver.Hom.unop f map_id' X := Opens.comap_id #align Top_op_to_Frame topOpToFrame diff --git a/Mathbin/Order/Category/HeytAlg.lean b/Mathbin/Order/Category/HeytAlg.lean index c02170da9d..24e9517c79 100644 --- a/Mathbin/Order/Category/HeytAlg.lean +++ b/Mathbin/Order/Category/HeytAlg.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies ! This file was ported from Lean 3 source module order.category.HeytAlg -! leanprover-community/mathlib commit 51cbe88849c5bcf9eaf8e38f2bbdf1a44bbabe0c +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.Order.Category.BoundedDistribLattice +import Mathbin.Order.Category.BddDistLat import Mathbin.Order.Heyting.Hom /-! @@ -59,11 +59,11 @@ instance bundledHom : BundledHom HeytingHom deriving instance LargeCategory, ConcreteCategory for HeytAlg @[simps] -instance hasForgetToLattice : HasForget₂ HeytAlg BoundedDistribLattice +instance hasForgetToLat : HasForget₂ HeytAlg BddDistLat where forget₂ := - { obj := fun X => BoundedDistribLattice.of X + { obj := fun X => BddDistLat.of X map := fun X Y f => (f : BoundedLatticeHom X Y) } -#align HeytAlg.has_forget_to_Lattice HeytAlg.hasForgetToLattice +#align HeytAlg.has_forget_to_Lat HeytAlg.hasForgetToLat /-- Constructs an isomorphism of Heyting algebras from an order isomorphism between them. -/ @[simps] diff --git a/Mathbin/Order/Category/Lattice.lean b/Mathbin/Order/Category/Lat.lean similarity index 50% rename from Mathbin/Order/Category/Lattice.lean rename to Mathbin/Order/Category/Lat.lean index 3343ddddba..3d912814c6 100644 --- a/Mathbin/Order/Category/Lattice.lean +++ b/Mathbin/Order/Category/Lat.lean @@ -3,26 +3,26 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -! This file was ported from Lean 3 source module order.category.Lattice -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 +! This file was ported from Lean 3 source module order.category.Lat +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.Order.Category.PartialOrder +import Mathbin.Order.Category.PartOrd import Mathbin.Order.Hom.Lattice /-! # The category of lattices -This defines `Lattice`, the category of lattices. +This defines `Lat`, the category of lattices. -Note that `Lattice` doesn't correspond to the literature definition of [`Lat`] +Note that `Lat` doesn't correspond to the literature definition of [`Lat`] (https://ncatlab.org/nlab/show/Lat) as we don't require bottom or top elements. Instead, `Lat` -corresponds to `BoundedLattice` (not yet in mathlib). +corresponds to `BddLat`. ## TODO -The free functor from `Lattice` to `BoundedLattice` is `X → with_top (with_bot X)`. +The free functor from `Lat` to `BddLat` is `X → with_top (with_bot X)`. -/ @@ -31,29 +31,29 @@ universe u open CategoryTheory /-- The category of lattices. -/ -def LatticeCat := +def Lat := Bundled Lattice -#align Lattice LatticeCat +#align Lat Lat -namespace LatticeCat +namespace Lat -instance : CoeSort LatticeCat (Type _) := +instance : CoeSort Lat (Type _) := Bundled.hasCoeToSort -instance (X : LatticeCat) : Lattice X := +instance (X : Lat) : Lattice X := X.str -/-- Construct a bundled `Lattice` from a `lattice`. -/ -def of (α : Type _) [Lattice α] : LatticeCat := +/-- Construct a bundled `Lat` from a `lattice`. -/ +def of (α : Type _) [Lattice α] : Lat := Bundled.of α -#align Lattice.of LatticeCat.of +#align Lat.of Lat.of @[simp] theorem coe_of (α : Type _) [Lattice α] : ↥(of α) = α := rfl -#align Lattice.coe_of LatticeCat.coe_of +#align Lat.coe_of Lat.coe_of -instance : Inhabited LatticeCat := +instance : Inhabited Lat := ⟨of Bool⟩ instance : BundledHom @LatticeHom where @@ -62,23 +62,23 @@ instance : BundledHom @LatticeHom where comp := @LatticeHom.comp hom_ext X Y _ _ := FunLike.coe_injective -instance : LargeCategory.{u} LatticeCat := +instance : LargeCategory.{u} Lat := BundledHom.category LatticeHom -instance : ConcreteCategory LatticeCat := +instance : ConcreteCategory Lat := BundledHom.concreteCategory LatticeHom -instance hasForgetToPartialOrder : HasForget₂ LatticeCat PartialOrderCat +instance hasForgetToPartOrd : HasForget₂ Lat PartOrd where forget₂ := { obj := fun X => ⟨X⟩ map := fun X Y f => f } forget_comp := rfl -#align Lattice.has_forget_to_PartialOrder LatticeCat.hasForgetToPartialOrder +#align Lat.has_forget_to_PartOrd Lat.hasForgetToPartOrd /-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ @[simps] -def Iso.mk {α β : LatticeCat.{u}} (e : α ≃o β) : α ≅ β +def Iso.mk {α β : Lat.{u}} (e : α ≃o β) : α ≅ β where Hom := e inv := e.symm @@ -88,28 +88,27 @@ def Iso.mk {α β : LatticeCat.{u}} (e : α ≃o β) : α ≅ β inv_hom_id' := by ext exact e.apply_symm_apply _ -#align Lattice.iso.mk LatticeCat.Iso.mk +#align Lat.iso.mk Lat.Iso.mk /-- `order_dual` as a functor. -/ @[simps] -def dual : LatticeCat ⥤ LatticeCat where +def dual : Lat ⥤ Lat where obj X := of Xᵒᵈ map X Y := LatticeHom.dual -#align Lattice.dual LatticeCat.dual +#align Lat.dual Lat.dual -/-- The equivalence between `Lattice` and itself induced by `order_dual` both ways. -/ +/-- The equivalence between `Lat` and itself induced by `order_dual` both ways. -/ @[simps Functor inverse] -def dualEquiv : LatticeCat ≌ LatticeCat := +def dualEquiv : Lat ≌ Lat := 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 Lattice.dual_equiv LatticeCat.dualEquiv +#align Lat.dual_equiv Lat.dualEquiv -end LatticeCat +end Lat -theorem latticeCat_dual_comp_forget_to_partialOrderCat : - LatticeCat.dual ⋙ forget₂ LatticeCat PartialOrderCat = - forget₂ LatticeCat PartialOrderCat ⋙ PartialOrderCat.dual := +theorem lat_dual_comp_forget_to_partOrd : + Lat.dual ⋙ forget₂ Lat PartOrd = forget₂ Lat PartOrd ⋙ PartOrd.dual := rfl -#align Lattice_dual_comp_forget_to_PartialOrder latticeCat_dual_comp_forget_to_partialOrderCat +#align Lat_dual_comp_forget_to_PartOrd lat_dual_comp_forget_to_partOrd diff --git a/Mathbin/Order/Category/LinOrd.lean b/Mathbin/Order/Category/LinOrd.lean new file mode 100644 index 0000000000..3fc967c3d7 --- /dev/null +++ b/Mathbin/Order/Category/LinOrd.lean @@ -0,0 +1,96 @@ +/- +Copyright (c) 2020 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin + +! This file was ported from Lean 3 source module order.category.LinOrd +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Order.Category.Lat + +/-! +# Category of linear orders + +This defines `LinOrd`, the category of linear orders with monotone maps. +-/ + + +open CategoryTheory + +universe u + +/-- The category of linear orders. -/ +def LinOrd := + Bundled LinearOrder +#align LinOrd LinOrd + +namespace LinOrd + +instance : BundledHom.ParentProjection @LinearOrder.toPartialOrder := + ⟨⟩ + +deriving instance LargeCategory, ConcreteCategory for LinOrd + +instance : CoeSort LinOrd (Type _) := + Bundled.hasCoeToSort + +/-- Construct a bundled `LinOrd` from the underlying type and typeclass. -/ +def of (α : Type _) [LinearOrder α] : LinOrd := + Bundled.of α +#align LinOrd.of LinOrd.of + +@[simp] +theorem coe_of (α : Type _) [LinearOrder α] : ↥(of α) = α := + rfl +#align LinOrd.coe_of LinOrd.coe_of + +instance : Inhabited LinOrd := + ⟨of PUnit⟩ + +instance (α : LinOrd) : LinearOrder α := + α.str + +instance hasForgetToLat : HasForget₂ LinOrd Lat + where forget₂ := + { obj := fun X => Lat.of X + map := fun X Y f => (OrderHomClass.toLatticeHom X Y f : LatticeHom X Y) } +#align LinOrd.has_forget_to_Lat LinOrd.hasForgetToLat + +/-- Constructs an equivalence between linear orders from an order isomorphism between them. -/ +@[simps] +def Iso.mk {α β : LinOrd.{u}} (e : α ≃o β) : α ≅ β + where + Hom := e + inv := e.symm + hom_inv_id' := by + ext + exact e.symm_apply_apply x + inv_hom_id' := by + ext + exact e.apply_symm_apply x +#align LinOrd.iso.mk LinOrd.Iso.mk + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : LinOrd ⥤ LinOrd where + obj X := of Xᵒᵈ + map X Y := OrderHom.dual +#align LinOrd.dual LinOrd.dual + +/-- The equivalence between `LinOrd` and itself induced by `order_dual` both ways. -/ +@[simps Functor inverse] +def dualEquiv : LinOrd ≌ LinOrd := + 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 LinOrd.dual_equiv LinOrd.dualEquiv + +end LinOrd + +theorem linOrd_dual_comp_forget_to_lat : + LinOrd.dual ⋙ forget₂ LinOrd Lat = forget₂ LinOrd Lat ⋙ Lat.dual := + rfl +#align LinOrd_dual_comp_forget_to_Lat linOrd_dual_comp_forget_to_lat + diff --git a/Mathbin/Order/Category/LinearOrder.lean b/Mathbin/Order/Category/LinearOrder.lean deleted file mode 100644 index 9ac8160387..0000000000 --- a/Mathbin/Order/Category/LinearOrder.lean +++ /dev/null @@ -1,98 +0,0 @@ -/- -Copyright (c) 2020 Johan Commelin. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johan Commelin - -! This file was ported from Lean 3 source module order.category.LinearOrder -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.Order.Category.Lattice - -/-! -# Category of linear orders - -This defines `LinearOrder`, the category of linear orders with monotone maps. --/ - - -open CategoryTheory - -universe u - -/-- The category of linear orders. -/ -def LinearOrderCat := - Bundled LinearOrder -#align LinearOrder LinearOrderCat - -namespace LinearOrderCat - -instance : BundledHom.ParentProjection @LinearOrder.toPartialOrder := - ⟨⟩ - -deriving instance LargeCategory, ConcreteCategory for LinearOrderCat - -instance : CoeSort LinearOrderCat (Type _) := - Bundled.hasCoeToSort - -/-- Construct a bundled `LinearOrder` from the underlying type and typeclass. -/ -def of (α : Type _) [LinearOrder α] : LinearOrderCat := - Bundled.of α -#align LinearOrder.of LinearOrderCat.of - -@[simp] -theorem coe_of (α : Type _) [LinearOrder α] : ↥(of α) = α := - rfl -#align LinearOrder.coe_of LinearOrderCat.coe_of - -instance : Inhabited LinearOrderCat := - ⟨of PUnit⟩ - -instance (α : LinearOrderCat) : LinearOrder α := - α.str - -instance hasForgetToLattice : HasForget₂ LinearOrderCat LatticeCat - where forget₂ := - { obj := fun X => LatticeCat.of X - map := fun X Y f => (OrderHomClass.toLatticeHom X Y f : LatticeHom X Y) } -#align LinearOrder.has_forget_to_Lattice LinearOrderCat.hasForgetToLattice - -/-- Constructs an equivalence between linear orders from an order isomorphism between them. -/ -@[simps] -def Iso.mk {α β : LinearOrderCat.{u}} (e : α ≃o β) : α ≅ β - where - Hom := e - inv := e.symm - hom_inv_id' := by - ext - exact e.symm_apply_apply x - inv_hom_id' := by - ext - exact e.apply_symm_apply x -#align LinearOrder.iso.mk LinearOrderCat.Iso.mk - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : LinearOrderCat ⥤ LinearOrderCat - where - obj X := of Xᵒᵈ - map X Y := OrderHom.dual -#align LinearOrder.dual LinearOrderCat.dual - -/-- The equivalence between `LinearOrder` and itself induced by `order_dual` both ways. -/ -@[simps Functor inverse] -def dualEquiv : LinearOrderCat ≌ LinearOrderCat := - 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 LinearOrder.dual_equiv LinearOrderCat.dualEquiv - -end LinearOrderCat - -theorem linearOrderCat_dual_comp_forget_to_latticeCat : - LinearOrderCat.dual ⋙ forget₂ LinearOrderCat LatticeCat = - forget₂ LinearOrderCat LatticeCat ⋙ LatticeCat.dual := - rfl -#align LinearOrder_dual_comp_forget_to_Lattice linearOrderCat_dual_comp_forget_to_latticeCat - diff --git a/Mathbin/Order/Category/NonemptyFinLinOrd.lean b/Mathbin/Order/Category/NonemptyFinLinOrd.lean index e647516e0d..3180a408ca 100644 --- a/Mathbin/Order/Category/NonemptyFinLinOrd.lean +++ b/Mathbin/Order/Category/NonemptyFinLinOrd.lean @@ -4,13 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin ! This file was ported from Lean 3 source module order.category.NonemptyFinLinOrd -! leanprover-community/mathlib commit 0bd2ea37bcba5769e14866170f251c9bc64e35d7 +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ import Mathbin.Data.Fintype.Order import Mathbin.Data.Set.Finite -import Mathbin.Order.Category.LinearOrder +import Mathbin.Order.Category.LinOrd import Mathbin.CategoryTheory.Limits.Shapes.Images import Mathbin.CategoryTheory.Limits.Shapes.RegularMono @@ -83,9 +83,9 @@ instance : Inhabited NonemptyFinLinOrdCat := instance (α : NonemptyFinLinOrdCat) : NonemptyFinLinOrd α := α.str -instance hasForgetToLinearOrder : HasForget₂ NonemptyFinLinOrdCat LinearOrderCat := +instance hasForgetToLinOrd : HasForget₂ NonemptyFinLinOrdCat LinOrd := BundledHom.forget₂ _ _ -#align NonemptyFinLinOrd.has_forget_to_LinearOrder NonemptyFinLinOrdCat.hasForgetToLinearOrder +#align NonemptyFinLinOrd.has_forget_to_LinOrd NonemptyFinLinOrdCat.hasForgetToLinOrd /-- Constructs an equivalence between nonempty finite linear orders from an order isomorphism between them. -/ @@ -110,7 +110,7 @@ def dual : NonemptyFinLinOrdCat ⥤ NonemptyFinLinOrdCat map X Y := OrderHom.dual #align NonemptyFinLinOrd.dual NonemptyFinLinOrdCat.dual -/-- The equivalence between `FinPartialOrder` and itself induced by `order_dual` both ways. -/ +/-- The equivalence between `FinPartOrd` and itself induced by `order_dual` both ways. -/ @[simps Functor inverse] def dualEquiv : NonemptyFinLinOrdCat ≌ NonemptyFinLinOrdCat := Equivalence.mk dual dual @@ -220,9 +220,9 @@ instance : HasStrongEpiMonoFactorisations NonemptyFinLinOrdCat.{u} := end NonemptyFinLinOrdCat -theorem nonemptyFinLinOrdCat_dual_comp_forget_to_linearOrderCat : - NonemptyFinLinOrdCat.dual ⋙ forget₂ NonemptyFinLinOrdCat LinearOrderCat = - forget₂ NonemptyFinLinOrdCat LinearOrderCat ⋙ LinearOrderCat.dual := +theorem nonemptyFinLinOrdCat_dual_comp_forget_to_linOrd : + NonemptyFinLinOrdCat.dual ⋙ forget₂ NonemptyFinLinOrdCat LinOrd = + forget₂ NonemptyFinLinOrdCat LinOrd ⋙ LinOrd.dual := rfl -#align NonemptyFinLinOrd_dual_comp_forget_to_LinearOrder nonemptyFinLinOrdCat_dual_comp_forget_to_linearOrderCat +#align NonemptyFinLinOrd_dual_comp_forget_to_LinOrd nonemptyFinLinOrdCat_dual_comp_forget_to_linOrd diff --git a/Mathbin/Order/Category/PartOrd.lean b/Mathbin/Order/Category/PartOrd.lean new file mode 100644 index 0000000000..766877609f --- /dev/null +++ b/Mathbin/Order/Category/PartOrd.lean @@ -0,0 +1,134 @@ +/- +Copyright (c) 2020 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin + +! This file was ported from Lean 3 source module order.category.PartOrd +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Order.Antisymmetrization +import Mathbin.Order.Category.Preord + +/-! +# Category of partial orders + +This defines `PartOrd`, the category of partial orders with monotone maps. +-/ + + +open CategoryTheory + +universe u + +/-- The category of partially ordered types. -/ +def PartOrd := + Bundled PartialOrder +#align PartOrd PartOrd + +namespace PartOrd + +instance : BundledHom.ParentProjection @PartialOrder.toPreorder := + ⟨⟩ + +deriving instance LargeCategory, ConcreteCategory for PartOrd + +instance : CoeSort PartOrd (Type _) := + Bundled.hasCoeToSort + +/-- Construct a bundled PartOrd from the underlying type and typeclass. -/ +def of (α : Type _) [PartialOrder α] : PartOrd := + Bundled.of α +#align PartOrd.of PartOrd.of + +@[simp] +theorem coe_of (α : Type _) [PartialOrder α] : ↥(of α) = α := + rfl +#align PartOrd.coe_of PartOrd.coe_of + +instance : Inhabited PartOrd := + ⟨of PUnit⟩ + +instance (α : PartOrd) : PartialOrder α := + α.str + +instance hasForgetToPreord : HasForget₂ PartOrd Preord := + BundledHom.forget₂ _ _ +#align PartOrd.has_forget_to_Preord PartOrd.hasForgetToPreord + +/-- Constructs an equivalence between partial orders from an order isomorphism between them. -/ +@[simps] +def Iso.mk {α β : PartOrd.{u}} (e : α ≃o β) : α ≅ β + where + Hom := e + inv := e.symm + hom_inv_id' := by + ext + exact e.symm_apply_apply x + inv_hom_id' := by + ext + exact e.apply_symm_apply x +#align PartOrd.iso.mk PartOrd.Iso.mk + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : PartOrd ⥤ PartOrd where + obj X := of Xᵒᵈ + map X Y := OrderHom.dual +#align PartOrd.dual PartOrd.dual + +/-- The equivalence between `PartOrd` and itself induced by `order_dual` both ways. -/ +@[simps Functor inverse] +def dualEquiv : PartOrd ≌ PartOrd := + 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 PartOrd.dual_equiv PartOrd.dualEquiv + +end PartOrd + +theorem partOrd_dual_comp_forget_to_preord : + PartOrd.dual ⋙ forget₂ PartOrd Preord = forget₂ PartOrd Preord ⋙ Preord.dual := + rfl +#align PartOrd_dual_comp_forget_to_Preord partOrd_dual_comp_forget_to_preord + +/-- `antisymmetrization` as a functor. It is the free functor. -/ +def preordToPartOrd : Preord.{u} ⥤ PartOrd + where + obj X := PartOrd.of (Antisymmetrization X (· ≤ ·)) + map X Y f := f.Antisymmetrization + map_id' X := by + ext + exact Quotient.inductionOn' x fun x => Quotient.map'_mk'' _ (fun a b => id) _ + map_comp' X Y Z f g := by + ext + exact Quotient.inductionOn' x fun x => OrderHom.antisymmetrization_apply_mk _ _ +#align Preord_to_PartOrd preordToPartOrd + +/-- `Preord_to_PartOrd` is left adjoint to the forgetful functor, meaning it is the free +functor from `Preord` to `PartOrd`. -/ +def preordToPartOrdForgetAdjunction : preordToPartOrd.{u} ⊣ forget₂ PartOrd Preord := + Adjunction.mkOfHomEquiv + { homEquiv := fun X Y => + { toFun := fun f => + ⟨f ∘ toAntisymmetrization (· ≤ ·), f.mono.comp toAntisymmetrization_mono⟩ + invFun := fun f => + ⟨fun a => Quotient.liftOn' a f fun a b h => (AntisymmRel.image h f.mono).Eq, fun a b => + Quotient.inductionOn₂' a b fun a b h => f.mono h⟩ + left_inv := fun f => + OrderHom.ext _ _ <| funext fun x => Quotient.inductionOn' x fun x => rfl + right_inv := fun f => OrderHom.ext _ _ <| funext fun x => rfl } + homEquiv_naturality_left_symm := fun X Y Z f g => + OrderHom.ext _ _ <| funext fun x => Quotient.inductionOn' x fun x => rfl + homEquiv_naturality_right := fun X Y Z f g => OrderHom.ext _ _ <| funext fun x => rfl } +#align Preord_to_PartOrd_forget_adjunction preordToPartOrdForgetAdjunction + +/-- `Preord_to_PartOrd` and `order_dual` commute. -/ +@[simps] +def preordToPartOrdCompToDualIsoToDualCompPreordToPartOrd : + preordToPartOrd.{u} ⋙ PartOrd.dual ≅ Preord.dual ⋙ preordToPartOrd := + NatIso.ofComponents (fun X => PartOrd.Iso.mk <| OrderIso.dualAntisymmetrization _) fun X Y f => + OrderHom.ext _ _ <| funext fun x => Quotient.inductionOn' x fun x => rfl +#align Preord_to_PartOrd_comp_to_dual_iso_to_dual_comp_Preord_to_PartOrd preordToPartOrdCompToDualIsoToDualCompPreordToPartOrd + diff --git a/Mathbin/Order/Category/PartialOrder.lean b/Mathbin/Order/Category/PartialOrder.lean deleted file mode 100644 index 4ead5fcf26..0000000000 --- a/Mathbin/Order/Category/PartialOrder.lean +++ /dev/null @@ -1,137 +0,0 @@ -/- -Copyright (c) 2020 Johan Commelin. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johan Commelin - -! This file was ported from Lean 3 source module order.category.PartialOrder -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.Order.Antisymmetrization -import Mathbin.Order.Category.Preorder - -/-! -# Category of partial orders - -This defines `PartialOrder`, the category of partial orders with monotone maps. --/ - - -open CategoryTheory - -universe u - -/-- The category of partially ordered types. -/ -def PartialOrderCat := - Bundled PartialOrder -#align PartialOrder PartialOrderCat - -namespace PartialOrderCat - -instance : BundledHom.ParentProjection @PartialOrder.toPreorder := - ⟨⟩ - -deriving instance LargeCategory, ConcreteCategory for PartialOrderCat - -instance : CoeSort PartialOrderCat (Type _) := - Bundled.hasCoeToSort - -/-- Construct a bundled PartialOrder from the underlying type and typeclass. -/ -def of (α : Type _) [PartialOrder α] : PartialOrderCat := - Bundled.of α -#align PartialOrder.of PartialOrderCat.of - -@[simp] -theorem coe_of (α : Type _) [PartialOrder α] : ↥(of α) = α := - rfl -#align PartialOrder.coe_of PartialOrderCat.coe_of - -instance : Inhabited PartialOrderCat := - ⟨of PUnit⟩ - -instance (α : PartialOrderCat) : PartialOrder α := - α.str - -instance hasForgetToPreorder : HasForget₂ PartialOrderCat PreorderCat := - BundledHom.forget₂ _ _ -#align PartialOrder.has_forget_to_Preorder PartialOrderCat.hasForgetToPreorder - -/-- Constructs an equivalence between partial orders from an order isomorphism between them. -/ -@[simps] -def Iso.mk {α β : PartialOrderCat.{u}} (e : α ≃o β) : α ≅ β - where - Hom := e - inv := e.symm - hom_inv_id' := by - ext - exact e.symm_apply_apply x - inv_hom_id' := by - ext - exact e.apply_symm_apply x -#align PartialOrder.iso.mk PartialOrderCat.Iso.mk - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : PartialOrderCat ⥤ PartialOrderCat - where - obj X := of Xᵒᵈ - map X Y := OrderHom.dual -#align PartialOrder.dual PartialOrderCat.dual - -/-- The equivalence between `PartialOrder` and itself induced by `order_dual` both ways. -/ -@[simps Functor inverse] -def dualEquiv : PartialOrderCat ≌ PartialOrderCat := - 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 PartialOrder.dual_equiv PartialOrderCat.dualEquiv - -end PartialOrderCat - -theorem partialOrderCat_dual_comp_forget_to_preorderCat : - PartialOrderCat.dual ⋙ forget₂ PartialOrderCat PreorderCat = - forget₂ PartialOrderCat PreorderCat ⋙ PreorderCat.dual := - rfl -#align PartialOrder_dual_comp_forget_to_Preorder partialOrderCat_dual_comp_forget_to_preorderCat - -/-- `antisymmetrization` as a functor. It is the free functor. -/ -def preorderToPartialOrder : PreorderCat.{u} ⥤ PartialOrderCat - where - obj X := PartialOrderCat.of (Antisymmetrization X (· ≤ ·)) - map X Y f := f.Antisymmetrization - map_id' X := by - ext - exact Quotient.inductionOn' x fun x => Quotient.map'_mk'' _ (fun a b => id) _ - map_comp' X Y Z f g := by - ext - exact Quotient.inductionOn' x fun x => OrderHom.antisymmetrization_apply_mk _ _ -#align Preorder_to_PartialOrder preorderToPartialOrder - -/-- `Preorder_to_PartialOrder` is left adjoint to the forgetful functor, meaning it is the free -functor from `Preorder` to `PartialOrder`. -/ -def preorderToPartialOrderForgetAdjunction : - preorderToPartialOrder.{u} ⊣ forget₂ PartialOrderCat PreorderCat := - Adjunction.mkOfHomEquiv - { homEquiv := fun X Y => - { toFun := fun f => - ⟨f ∘ toAntisymmetrization (· ≤ ·), f.mono.comp toAntisymmetrization_mono⟩ - invFun := fun f => - ⟨fun a => Quotient.liftOn' a f fun a b h => (AntisymmRel.image h f.mono).Eq, fun a b => - Quotient.inductionOn₂' a b fun a b h => f.mono h⟩ - left_inv := fun f => - OrderHom.ext _ _ <| funext fun x => Quotient.inductionOn' x fun x => rfl - right_inv := fun f => OrderHom.ext _ _ <| funext fun x => rfl } - homEquiv_naturality_left_symm := fun X Y Z f g => - OrderHom.ext _ _ <| funext fun x => Quotient.inductionOn' x fun x => rfl - homEquiv_naturality_right := fun X Y Z f g => OrderHom.ext _ _ <| funext fun x => rfl } -#align Preorder_to_PartialOrder_forget_adjunction preorderToPartialOrderForgetAdjunction - -/-- `Preorder_to_PartialOrder` and `order_dual` commute. -/ -@[simps] -def preorderToPartialOrderCompToDualIsoToDualCompPreorderToPartialOrder : - preorderToPartialOrder.{u} ⋙ PartialOrderCat.dual ≅ PreorderCat.dual ⋙ preorderToPartialOrder := - NatIso.ofComponents (fun X => PartialOrderCat.Iso.mk <| OrderIso.dualAntisymmetrization _) - fun X Y f => OrderHom.ext _ _ <| funext fun x => Quotient.inductionOn' x fun x => rfl -#align Preorder_to_PartialOrder_comp_to_dual_iso_to_dual_comp_Preorder_to_PartialOrder preorderToPartialOrderCompToDualIsoToDualCompPreorderToPartialOrder - diff --git a/Mathbin/Order/Category/Preorder.lean b/Mathbin/Order/Category/Preord.lean similarity index 60% rename from Mathbin/Order/Category/Preorder.lean rename to Mathbin/Order/Category/Preord.lean index 8864ef8e56..db75225df3 100644 --- a/Mathbin/Order/Category/Preorder.lean +++ b/Mathbin/Order/Category/Preord.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -! This file was ported from Lean 3 source module order.category.Preorder -! leanprover-community/mathlib commit 5306f2df04c869c95c48367d75ad3051c359e9b5 +! This file was ported from Lean 3 source module order.category.Preord +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -16,7 +16,7 @@ import Mathbin.Order.Hom.Basic /-! # Category of preorders -This defines `Preorder`, the category of preorders with monotone maps. +This defines `Preord`, the category of preorders with monotone maps. -/ @@ -25,11 +25,11 @@ universe u open CategoryTheory /-- The category of preorders. -/ -def PreorderCat := +def Preord := Bundled Preorder -#align Preorder PreorderCat +#align Preord Preord -namespace PreorderCat +namespace Preord instance : BundledHom @OrderHom where toFun := @OrderHom.toFun @@ -37,30 +37,30 @@ instance : BundledHom @OrderHom where comp := @OrderHom.comp hom_ext := @OrderHom.ext -deriving instance LargeCategory, ConcreteCategory for PreorderCat +deriving instance LargeCategory, ConcreteCategory for Preord -instance : CoeSort PreorderCat (Type _) := +instance : CoeSort Preord (Type _) := Bundled.hasCoeToSort -/-- Construct a bundled Preorder from the underlying type and typeclass. -/ -def of (α : Type _) [Preorder α] : PreorderCat := +/-- Construct a bundled Preord from the underlying type and typeclass. -/ +def of (α : Type _) [Preorder α] : Preord := Bundled.of α -#align Preorder.of PreorderCat.of +#align Preord.of Preord.of @[simp] theorem coe_of (α : Type _) [Preorder α] : ↥(of α) = α := rfl -#align Preorder.coe_of PreorderCat.coe_of +#align Preord.coe_of Preord.coe_of -instance : Inhabited PreorderCat := +instance : Inhabited Preord := ⟨of PUnit⟩ -instance (α : PreorderCat) : Preorder α := +instance (α : Preord) : Preorder α := α.str /-- Constructs an equivalence between preorders from an order isomorphism between them. -/ @[simps] -def Iso.mk {α β : PreorderCat.{u}} (e : α ≃o β) : α ≅ β +def Iso.mk {α β : Preord.{u}} (e : α ≃o β) : α ≅ β where Hom := e inv := e.symm @@ -70,40 +70,39 @@ def Iso.mk {α β : PreorderCat.{u}} (e : α ≃o β) : α ≅ β inv_hom_id' := by ext exact e.apply_symm_apply x -#align Preorder.iso.mk PreorderCat.Iso.mk +#align Preord.iso.mk Preord.Iso.mk /-- `order_dual` as a functor. -/ @[simps] -def dual : PreorderCat ⥤ PreorderCat where +def dual : Preord ⥤ Preord where obj X := of Xᵒᵈ map X Y := OrderHom.dual -#align Preorder.dual PreorderCat.dual +#align Preord.dual Preord.dual -/-- The equivalence between `Preorder` and itself induced by `order_dual` both ways. -/ +/-- The equivalence between `Preord` and itself induced by `order_dual` both ways. -/ @[simps Functor inverse] -def dualEquiv : PreorderCat ≌ PreorderCat := +def dualEquiv : Preord ≌ Preord := 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 Preorder.dual_equiv PreorderCat.dualEquiv +#align Preord.dual_equiv Preord.dualEquiv -end PreorderCat +end Preord -/-- The embedding of `Preorder` into `Cat`. +/-- The embedding of `Preord` into `Cat`. -/ @[simps] -def preorderToCat : PreorderCat.{u} ⥤ Cat - where +def preordToCat : Preord.{u} ⥤ Cat where obj X := Cat.of X.1 map X Y f := f.Monotone.Functor map_id' X := by apply CategoryTheory.Functor.ext; tidy map_comp' X Y Z f g := by apply CategoryTheory.Functor.ext; tidy -#align Preorder_to_Cat preorderToCat +#align Preord_to_Cat preordToCat -instance : Faithful preorderToCat.{u} +instance : Faithful preordToCat.{u} where map_injective' X Y f g h := by ext x; exact functor.congr_obj h x -instance : Full preorderToCat.{u} +instance : Full preordToCat.{u} where preimage X Y f := ⟨f.obj, f.Monotone⟩ witness' X Y f := by apply CategoryTheory.Functor.ext; tidy diff --git a/Mathbin/Order/Category/Semilat.lean b/Mathbin/Order/Category/Semilat.lean new file mode 100644 index 0000000000..a3901a46ee --- /dev/null +++ b/Mathbin/Order/Category/Semilat.lean @@ -0,0 +1,217 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies + +! This file was ported from Lean 3 source module order.category.Semilat +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Order.Category.PartOrd +import Mathbin.Order.Hom.Lattice + +/-! +# The categories of semilattices + +This defines `SemilatSup` and `SemilatInf`, the categories of sup-semilattices with a bottom +element and inf-semilattices with a top element. + +## References + +* [nLab, *semilattice*](https://ncatlab.org/nlab/show/semilattice) +-/ + + +universe u + +open CategoryTheory + +/-- The category of sup-semilattices with a bottom element. -/ +structure SemilatSup : Type (u + 1) where + pt : Type u + [isSemilatticeSup : SemilatticeSup X] + [isOrderBot : OrderBot X] +#align SemilatSup SemilatSup + +/-- The category of inf-semilattices with a top element. -/ +structure SemilatInf : Type (u + 1) where + pt : Type u + [isSemilatticeInf : SemilatticeInf X] + [isOrderTop : OrderTop X] +#align SemilatInf SemilatInf + +attribute [protected] SemilatSup.X SemilatInf.X + +namespace SemilatSup + +instance : CoeSort SemilatSup (Type _) := + ⟨SemilatSup.X⟩ + +attribute [instance] is_semilattice_sup is_order_bot + +/-- Construct a bundled `SemilatSup` from a `semilattice_sup`. -/ +def of (α : Type _) [SemilatticeSup α] [OrderBot α] : SemilatSup := + ⟨α⟩ +#align SemilatSup.of SemilatSup.of + +@[simp] +theorem coe_of (α : Type _) [SemilatticeSup α] [OrderBot α] : ↥(of α) = α := + rfl +#align SemilatSup.coe_of SemilatSup.coe_of + +instance : Inhabited SemilatSup := + ⟨of PUnit⟩ + +instance : LargeCategory.{u} SemilatSup + where + Hom X Y := SupBotHom X Y + id X := SupBotHom.id X + comp X Y Z f g := g.comp f + id_comp' X Y := SupBotHom.comp_id + comp_id' X Y := SupBotHom.id_comp + assoc' W X Y Z _ _ _ := SupBotHom.comp_assoc _ _ _ + +instance : ConcreteCategory SemilatSup + where + forget := + { obj := SemilatSup.X + map := fun X Y => coeFn } + forget_faithful := ⟨fun X Y => FunLike.coe_injective⟩ + +instance hasForgetToPartOrd : HasForget₂ SemilatSup PartOrd + where forget₂ := + { obj := fun X => ⟨X⟩ + map := fun X Y f => f } +#align SemilatSup.has_forget_to_PartOrd SemilatSup.hasForgetToPartOrd + +@[simp] +theorem coe_forget_to_partOrd (X : SemilatSup) : ↥((forget₂ SemilatSup PartOrd).obj X) = ↥X := + rfl +#align SemilatSup.coe_forget_to_PartOrd SemilatSup.coe_forget_to_partOrd + +end SemilatSup + +namespace SemilatInf + +instance : CoeSort SemilatInf (Type _) := + ⟨SemilatInf.X⟩ + +attribute [instance] is_semilattice_inf is_order_top + +/-- Construct a bundled `SemilatInf` from a `semilattice_inf`. -/ +def of (α : Type _) [SemilatticeInf α] [OrderTop α] : SemilatInf := + ⟨α⟩ +#align SemilatInf.of SemilatInf.of + +@[simp] +theorem coe_of (α : Type _) [SemilatticeInf α] [OrderTop α] : ↥(of α) = α := + rfl +#align SemilatInf.coe_of SemilatInf.coe_of + +instance : Inhabited SemilatInf := + ⟨of PUnit⟩ + +instance : LargeCategory.{u} SemilatInf + where + Hom X Y := InfTopHom X Y + id X := InfTopHom.id X + comp X Y Z f g := g.comp f + id_comp' X Y := InfTopHom.comp_id + comp_id' X Y := InfTopHom.id_comp + assoc' W X Y Z _ _ _ := InfTopHom.comp_assoc _ _ _ + +instance : ConcreteCategory SemilatInf + where + forget := + { obj := SemilatInf.X + map := fun X Y => coeFn } + forget_faithful := ⟨fun X Y => FunLike.coe_injective⟩ + +instance hasForgetToPartOrd : HasForget₂ SemilatInf PartOrd + where forget₂ := + { obj := fun X => ⟨X⟩ + map := fun X Y f => f } +#align SemilatInf.has_forget_to_PartOrd SemilatInf.hasForgetToPartOrd + +@[simp] +theorem coe_forget_to_partOrd (X : SemilatInf) : ↥((forget₂ SemilatInf PartOrd).obj X) = ↥X := + rfl +#align SemilatInf.coe_forget_to_PartOrd SemilatInf.coe_forget_to_partOrd + +end SemilatInf + +/-! ### Order dual -/ + + +namespace SemilatSup + +/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ +@[simps] +def Iso.mk {α β : SemilatSup.{u}} (e : α ≃o β) : α ≅ β + where + Hom := e + inv := e.symm + hom_inv_id' := by + ext + exact e.symm_apply_apply _ + inv_hom_id' := by + ext + exact e.apply_symm_apply _ +#align SemilatSup.iso.mk SemilatSup.Iso.mk + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : SemilatSup ⥤ SemilatInf + where + obj X := SemilatInf.of Xᵒᵈ + map X Y := SupBotHom.dual +#align SemilatSup.dual SemilatSup.dual + +end SemilatSup + +namespace SemilatInf + +/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ +@[simps] +def Iso.mk {α β : SemilatInf.{u}} (e : α ≃o β) : α ≅ β + where + Hom := e + inv := e.symm + hom_inv_id' := by + ext + exact e.symm_apply_apply _ + inv_hom_id' := by + ext + exact e.apply_symm_apply _ +#align SemilatInf.iso.mk SemilatInf.Iso.mk + +/-- `order_dual` as a functor. -/ +@[simps] +def dual : SemilatInf ⥤ SemilatSup + where + obj X := SemilatSup.of Xᵒᵈ + map X Y := InfTopHom.dual +#align SemilatInf.dual SemilatInf.dual + +end SemilatInf + +/-- The equivalence between `SemilatSup` and `SemilatInf` induced by `order_dual` both ways. +-/ +@[simps Functor inverse] +def semilatSupEquivSemilatInf : SemilatSup ≌ SemilatInf := + Equivalence.mk SemilatSup.dual SemilatInf.dual + (NatIso.ofComponents (fun X => SemilatSup.Iso.mk <| OrderIso.dualDual X) fun X Y f => rfl) + (NatIso.ofComponents (fun X => SemilatInf.Iso.mk <| OrderIso.dualDual X) fun X Y f => rfl) +#align SemilatSup_equiv_SemilatInf semilatSupEquivSemilatInf + +theorem semilatSup_dual_comp_forget_to_partOrd : + SemilatSup.dual ⋙ forget₂ SemilatInf PartOrd = forget₂ SemilatSup PartOrd ⋙ PartOrd.dual := + rfl +#align SemilatSup_dual_comp_forget_to_PartOrd semilatSup_dual_comp_forget_to_partOrd + +theorem semilatInf_dual_comp_forget_to_partOrd : + SemilatInf.dual ⋙ forget₂ SemilatSup PartOrd = forget₂ SemilatInf PartOrd ⋙ PartOrd.dual := + rfl +#align SemilatInf_dual_comp_forget_to_PartOrd semilatInf_dual_comp_forget_to_partOrd + diff --git a/Mathbin/Order/Category/Semilattice.lean b/Mathbin/Order/Category/Semilattice.lean deleted file mode 100644 index 394a8fb0b5..0000000000 --- a/Mathbin/Order/Category/Semilattice.lean +++ /dev/null @@ -1,223 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies - -! This file was ported from Lean 3 source module order.category.Semilattice -! leanprover-community/mathlib commit a057441eb14fe0e65a5588084e7e2c1193b82579 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathbin.Order.Category.PartialOrder -import Mathbin.Order.Hom.Lattice - -/-! -# The categories of semilattices - -This defines `SemilatticeSup` and `SemilatticeInf`, the categories of sup-semilattices with a bottom -element and inf-semilattices with a top element. - -## References - -* [nLab, *semilattice*](https://ncatlab.org/nlab/show/semilattice) --/ - - -universe u - -open CategoryTheory - -/-- The category of sup-semilattices with a bottom element. -/ -structure SemilatticeSupCat : Type (u + 1) where - pt : Type u - [isSemilatticeSup : SemilatticeSup X] - [isOrderBot : OrderBot X] -#align SemilatticeSup SemilatticeSupCat - -/-- The category of inf-semilattices with a top element. -/ -structure SemilatticeInfCat : Type (u + 1) where - pt : Type u - [isSemilatticeInf : SemilatticeInf X] - [isOrderTop : OrderTop X] -#align SemilatticeInf SemilatticeInfCat - -attribute [protected] SemilatticeSupCat.X SemilatticeInfCat.X - -namespace SemilatticeSupCat - -instance : CoeSort SemilatticeSupCat (Type _) := - ⟨SemilatticeSupCat.X⟩ - -attribute [instance] is_semilattice_sup is_order_bot - -/-- Construct a bundled `SemilatticeSup` from a `semilattice_sup`. -/ -def of (α : Type _) [SemilatticeSup α] [OrderBot α] : SemilatticeSupCat := - ⟨α⟩ -#align SemilatticeSup.of SemilatticeSupCat.of - -@[simp] -theorem coe_of (α : Type _) [SemilatticeSup α] [OrderBot α] : ↥(of α) = α := - rfl -#align SemilatticeSup.coe_of SemilatticeSupCat.coe_of - -instance : Inhabited SemilatticeSupCat := - ⟨of PUnit⟩ - -instance : LargeCategory.{u} SemilatticeSupCat - where - Hom X Y := SupBotHom X Y - id X := SupBotHom.id X - comp X Y Z f g := g.comp f - id_comp' X Y := SupBotHom.comp_id - comp_id' X Y := SupBotHom.id_comp - assoc' W X Y Z _ _ _ := SupBotHom.comp_assoc _ _ _ - -instance : ConcreteCategory SemilatticeSupCat - where - forget := - { obj := SemilatticeSupCat.X - map := fun X Y => coeFn } - forget_faithful := ⟨fun X Y => FunLike.coe_injective⟩ - -instance hasForgetToPartialOrder : HasForget₂ SemilatticeSupCat PartialOrderCat - where forget₂ := - { obj := fun X => ⟨X⟩ - map := fun X Y f => f } -#align SemilatticeSup.has_forget_to_PartialOrder SemilatticeSupCat.hasForgetToPartialOrder - -@[simp] -theorem coe_forget_to_partialOrderCat (X : SemilatticeSupCat) : - ↥((forget₂ SemilatticeSupCat PartialOrderCat).obj X) = ↥X := - rfl -#align SemilatticeSup.coe_forget_to_PartialOrder SemilatticeSupCat.coe_forget_to_partialOrderCat - -end SemilatticeSupCat - -namespace SemilatticeInfCat - -instance : CoeSort SemilatticeInfCat (Type _) := - ⟨SemilatticeInfCat.X⟩ - -attribute [instance] is_semilattice_inf is_order_top - -/-- Construct a bundled `SemilatticeInf` from a `semilattice_inf`. -/ -def of (α : Type _) [SemilatticeInf α] [OrderTop α] : SemilatticeInfCat := - ⟨α⟩ -#align SemilatticeInf.of SemilatticeInfCat.of - -@[simp] -theorem coe_of (α : Type _) [SemilatticeInf α] [OrderTop α] : ↥(of α) = α := - rfl -#align SemilatticeInf.coe_of SemilatticeInfCat.coe_of - -instance : Inhabited SemilatticeInfCat := - ⟨of PUnit⟩ - -instance : LargeCategory.{u} SemilatticeInfCat - where - Hom X Y := InfTopHom X Y - id X := InfTopHom.id X - comp X Y Z f g := g.comp f - id_comp' X Y := InfTopHom.comp_id - comp_id' X Y := InfTopHom.id_comp - assoc' W X Y Z _ _ _ := InfTopHom.comp_assoc _ _ _ - -instance : ConcreteCategory SemilatticeInfCat - where - forget := - { obj := SemilatticeInfCat.X - map := fun X Y => coeFn } - forget_faithful := ⟨fun X Y => FunLike.coe_injective⟩ - -instance hasForgetToPartialOrder : HasForget₂ SemilatticeInfCat PartialOrderCat - where forget₂ := - { obj := fun X => ⟨X⟩ - map := fun X Y f => f } -#align SemilatticeInf.has_forget_to_PartialOrder SemilatticeInfCat.hasForgetToPartialOrder - -@[simp] -theorem coe_forget_to_partialOrderCat (X : SemilatticeInfCat) : - ↥((forget₂ SemilatticeInfCat PartialOrderCat).obj X) = ↥X := - rfl -#align SemilatticeInf.coe_forget_to_PartialOrder SemilatticeInfCat.coe_forget_to_partialOrderCat - -end SemilatticeInfCat - -/-! ### Order dual -/ - - -namespace SemilatticeSupCat - -/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ -@[simps] -def Iso.mk {α β : SemilatticeSupCat.{u}} (e : α ≃o β) : α ≅ β - where - Hom := e - inv := e.symm - hom_inv_id' := by - ext - exact e.symm_apply_apply _ - inv_hom_id' := by - ext - exact e.apply_symm_apply _ -#align SemilatticeSup.iso.mk SemilatticeSupCat.Iso.mk - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : SemilatticeSupCat ⥤ SemilatticeInfCat - where - obj X := SemilatticeInfCat.of Xᵒᵈ - map X Y := SupBotHom.dual -#align SemilatticeSup.dual SemilatticeSupCat.dual - -end SemilatticeSupCat - -namespace SemilatticeInfCat - -/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ -@[simps] -def Iso.mk {α β : SemilatticeInfCat.{u}} (e : α ≃o β) : α ≅ β - where - Hom := e - inv := e.symm - hom_inv_id' := by - ext - exact e.symm_apply_apply _ - inv_hom_id' := by - ext - exact e.apply_symm_apply _ -#align SemilatticeInf.iso.mk SemilatticeInfCat.Iso.mk - -/-- `order_dual` as a functor. -/ -@[simps] -def dual : SemilatticeInfCat ⥤ SemilatticeSupCat - where - obj X := SemilatticeSupCat.of Xᵒᵈ - map X Y := InfTopHom.dual -#align SemilatticeInf.dual SemilatticeInfCat.dual - -end SemilatticeInfCat - -/-- The equivalence between `SemilatticeSup` and `SemilatticeInf` induced by `order_dual` both ways. --/ -@[simps Functor inverse] -def semilatticeSupEquivSemilatticeInf : SemilatticeSupCat ≌ SemilatticeInfCat := - Equivalence.mk SemilatticeSupCat.dual SemilatticeInfCat.dual - (NatIso.ofComponents (fun X => SemilatticeSupCat.Iso.mk <| OrderIso.dualDual X) fun X Y f => - rfl) - (NatIso.ofComponents (fun X => SemilatticeInfCat.Iso.mk <| OrderIso.dualDual X) fun X Y f => - rfl) -#align SemilatticeSup_equiv_SemilatticeInf semilatticeSupEquivSemilatticeInf - -theorem semilatticeSupCat_dual_comp_forget_to_partialOrderCat : - SemilatticeSupCat.dual ⋙ forget₂ SemilatticeInfCat PartialOrderCat = - forget₂ SemilatticeSupCat PartialOrderCat ⋙ PartialOrderCat.dual := - rfl -#align SemilatticeSup_dual_comp_forget_to_PartialOrder semilatticeSupCat_dual_comp_forget_to_partialOrderCat - -theorem semilatticeInfCat_dual_comp_forget_to_partialOrderCat : - SemilatticeInfCat.dual ⋙ forget₂ SemilatticeSupCat PartialOrderCat = - forget₂ SemilatticeInfCat PartialOrderCat ⋙ PartialOrderCat.dual := - rfl -#align SemilatticeInf_dual_comp_forget_to_PartialOrder semilatticeInfCat_dual_comp_forget_to_partialOrderCat - diff --git a/Mathbin/Order/CompleteBooleanAlgebra.lean b/Mathbin/Order/CompleteBooleanAlgebra.lean index 34621cec9b..dcd6833863 100644 --- a/Mathbin/Order/CompleteBooleanAlgebra.lean +++ b/Mathbin/Order/CompleteBooleanAlgebra.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yaël Dillies ! This file was ported from Lean 3 source module order.complete_boolean_algebra -! leanprover-community/mathlib commit c3291da49cfa65f0d43b094750541c0731edc932 +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -23,7 +23,7 @@ distributive Boolean algebras. ## Typeclasses -* `order.frame`: Frame: A complete lattice whose `⊓` distributes over `⨆`. +* `order.frame`: Frm: A complete lattice whose `⊓` distributes over `⨆`. * `order.coframe`: Coframe: A complete lattice whose `⊔` distributes over `⨅`. * `complete_distrib_lattice`: Completely distributive lattices: A complete lattice whose `⊓` and `⊔` distribute over `⨆` and `⨅` respectively. diff --git a/Mathbin/Order/Hom/CompleteLattice.lean b/Mathbin/Order/Hom/CompleteLattice.lean index d6452098e0..f8a7585f35 100644 --- a/Mathbin/Order/Hom/CompleteLattice.lean +++ b/Mathbin/Order/Hom/CompleteLattice.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies ! This file was ported from Lean 3 source module order.hom.complete_lattice -! leanprover-community/mathlib commit 50832daea47b195a48b5b33b1c8b2162c48c3afc +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -26,7 +26,7 @@ be satisfied by itself and all stricter types. * `Sup_hom`: Maps which preserve `⨆`. * `Inf_hom`: Maps which preserve `⨅`. -* `frame_hom`: Frame homomorphisms. Maps which preserve `⨆`, `⊓` and `⊤`. +* `frame_hom`: Frm homomorphisms. Maps which preserve `⨆`, `⊓` and `⊤`. * `complete_lattice_hom`: Complete lattice homomorphisms. Maps which preserve `⨆` and `⨅`. ## Typeclasses @@ -42,7 +42,7 @@ be satisfied by itself and all stricter types. ## TODO -Frame homs are Heyting homs. +Frm homs are Heyting homs. -/ @@ -782,7 +782,7 @@ theorem top_apply (a : α) : (⊤ : InfₛHom α β) a = ⊤ := end InfₛHom -/-! ### Frame homomorphisms -/ +/-! ### Frm homomorphisms -/ namespace FrameHom diff --git a/Mathbin/Order/UpperLower/LocallyFinite.lean b/Mathbin/Order/UpperLower/LocallyFinite.lean new file mode 100644 index 0000000000..8550bb9cd6 --- /dev/null +++ b/Mathbin/Order/UpperLower/LocallyFinite.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2023 Yaël Dillies, Sara Rousta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies + +! This file was ported from Lean 3 source module order.upper_lower.locally_finite +! leanprover-community/mathlib commit 3e175454d8aa6a94734afcb2ae20a2f2b6660ea5 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Order.LocallyFinite +import Mathbin.Order.UpperLower.Basic + +/-! +# Upper and lower sets in a locally finite order + +In this file we characterise the interaction of `upper_set`/`lower_set` and `locally_finite_order`. +-/ + + +namespace Set + +variable {α : Type _} [Preorder α] {s : Set α} + +protected theorem Finite.upperClosure [LocallyFiniteOrderTop α] (hs : s.Finite) : + (upperClosure s : Set α).Finite := by + rw [coe_upperClosure] + exact hs.bUnion fun _ _ => finite_Ici _ +#align set.finite.upper_closure Set.Finite.upperClosure + +protected theorem Finite.lowerClosure [LocallyFiniteOrderBot α] (hs : s.Finite) : + (lowerClosure s : Set α).Finite := by + rw [coe_lowerClosure] + exact hs.bUnion fun _ _ => finite_Iic _ +#align set.finite.lower_closure Set.Finite.lowerClosure + +end Set + diff --git a/Mathbin/RepresentationTheory/Basic.lean b/Mathbin/RepresentationTheory/Basic.lean index 91c3a5475b..98b9f5620a 100644 --- a/Mathbin/RepresentationTheory/Basic.lean +++ b/Mathbin/RepresentationTheory/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Labelle ! This file was ported from Lean 3 source module representation_theory.basic -! leanprover-community/mathlib commit 3dec44d0b621a174c56e994da4aae15ba60110a2 +! leanprover-community/mathlib commit 9cb7206107eb5cbc8dd3b9fc0864c548fae962bd ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -436,6 +436,12 @@ theorem dual_apply (g : G) : (dual ρV) g = Module.Dual.transpose (ρV g⁻¹) : rfl #align representation.dual_apply Representation.dual_apply +/-- Given $k$-modules $V, W$, there is a homomorphism $φ : V^* ⊗ W → Hom_k(V, W)$ +(implemented by `linear_algebra.contraction.dual_tensor_hom`). +Given representations of $G$ on $V$ and $W$,there are representations of $G$ on $V^* ⊗ W$ and on +$Hom_k(V, W)$. +This lemma says that $φ$ is $G$-linear. +-/ theorem dualTensorHom_comm (g : G) : dualTensorHom k V W ∘ₗ TensorProduct.map (ρV.dual g) (ρW g) = (linHom ρV ρW) g ∘ₗ dualTensorHom k V W := diff --git a/Mathbin/RingTheory/Ideal/Norm.lean b/Mathbin/RingTheory/Ideal/Norm.lean index a2bcdf64c9..c942228ea3 100644 --- a/Mathbin/RingTheory/Ideal/Norm.lean +++ b/Mathbin/RingTheory/Ideal/Norm.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Alex J. Best ! This file was ported from Lean 3 source module ring_theory.ideal.norm -! leanprover-community/mathlib commit 85e3c05a94b27c84dc6f234cf88326d5e0096ec3 +! leanprover-community/mathlib commit d3acee0d776b15ffb8318f327325ff343cc8bdcc ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -293,6 +293,11 @@ theorem absNorm_eq_one_iff {I : Ideal S} : absNorm I = 1 ↔ I = ⊤ := by rw [abs_norm_apply, card_quot_eq_one_iff] #align ideal.abs_norm_eq_one_iff Ideal.absNorm_eq_one_iff +theorem absNorm_ne_zero_iff (I : Ideal S) : Ideal.absNorm I ≠ 0 ↔ Finite (S ⧸ I) := + ⟨fun h => Nat.finite_of_card_ne_zero h, fun h => + (@AddSubgroup.finiteIndex_of_finite_quotient _ _ _ h).FiniteIndex⟩ +#align ideal.abs_norm_ne_zero_iff Ideal.absNorm_ne_zero_iff + /-- Let `e : S ≃ I` be an additive isomorphism (therefore a `ℤ`-linear equiv). Then an alternative way to compute the norm of `I` is given by taking the determinant of `e`. See `nat_abs_det_basis_change` for a more familiar formulation of this result. -/ @@ -460,6 +465,29 @@ theorem absNorm_mem (I : Ideal S) : ↑I.absNorm ∈ I := by quotient.index_eq_zero] #align ideal.abs_norm_mem Ideal.absNorm_mem +theorem span_singleton_absNorm_le (I : Ideal S) : Ideal.span {(Ideal.absNorm I : S)} ≤ I := by + simp only [Ideal.span_le, Set.singleton_subset_iff, SetLike.mem_coe, Ideal.absNorm_mem I] +#align ideal.span_singleton_abs_norm_le Ideal.span_singleton_absNorm_le + +theorem finite_setOf_absNorm_eq [CharZero S] {n : ℕ} (hn : 0 < n) : + { I : Ideal S | Ideal.absNorm I = n }.Finite := + by + let f := fun I : Ideal S => Ideal.map (Ideal.Quotient.mk (@Ideal.span S _ {n})) I + refine' @Set.Finite.of_finite_image _ _ _ f _ _ + · suffices Finite (S ⧸ @Ideal.span S _ {n}) + by + let g := (coe : Ideal (S ⧸ @Ideal.span S _ {n}) → Set (S ⧸ @Ideal.span S _ {n})) + refine' @Set.Finite.of_finite_image _ _ _ g _ (set_like.coe_injective.inj_on _) + exact Set.Finite.subset (@Set.finite_univ _ (@Set.finite' _ this)) (Set.subset_univ _) + rw [← abs_norm_ne_zero_iff, abs_norm_span_singleton] + simpa only [Ne.def, Int.natAbs_eq_zero, Algebra.norm_eq_zero_iff, Nat.cast_eq_zero] using + ne_of_gt hn + · intro I hI J hJ h + rw [← comap_map_mk (span_singleton_abs_norm_le I), ← hI.symm, ← + comap_map_mk (span_singleton_abs_norm_le J), ← hJ.symm] + exact congr_arg (Ideal.comap (Ideal.Quotient.mk (@Ideal.span S _ {n}))) h +#align ideal.finite_set_of_abs_norm_eq Ideal.finite_setOf_absNorm_eq + end Ideal end RingOfIntegers diff --git a/Mathbin/RingTheory/Ideal/QuotientOperations.lean b/Mathbin/RingTheory/Ideal/QuotientOperations.lean index 555e9693d7..7ca1f85228 100644 --- a/Mathbin/RingTheory/Ideal/QuotientOperations.lean +++ b/Mathbin/RingTheory/Ideal/QuotientOperations.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau ! This file was ported from Lean 3 source module ring_theory.ideal.quotient_operations -! leanprover-community/mathlib commit e7f0ddbf65bd7181a85edb74b64bdc35ba4bdc74 +! leanprover-community/mathlib commit d3acee0d776b15ffb8318f327325ff343cc8bdcc ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -150,6 +150,13 @@ theorem mem_quotient_iff_mem {I J : Ideal R} (hIJ : I ≤ J) {x : R} : rw [mem_quotient_iff_mem_sup, sup_eq_left.mpr hIJ] #align ideal.mem_quotient_iff_mem Ideal.mem_quotient_iff_mem +theorem comap_map_mk {I J : Ideal R} (h : I ≤ J) : + Ideal.comap (Ideal.Quotient.mk I) (Ideal.map (Ideal.Quotient.mk I) J) = J := + by + ext + rw [← Ideal.mem_quotient_iff_mem h, Ideal.mem_comap] +#align ideal.comap_map_mk Ideal.comap_map_mk + section QuotientAlgebra variable (R₁ R₂ : Type _) {A B : Type _} diff --git a/Mathbin/RingTheory/MvPolynomial/WeightedHomogeneous.lean b/Mathbin/RingTheory/MvPolynomial/WeightedHomogeneous.lean new file mode 100644 index 0000000000..16bc4e26a4 --- /dev/null +++ b/Mathbin/RingTheory/MvPolynomial/WeightedHomogeneous.lean @@ -0,0 +1,481 @@ +/- +Copyright (c) 2022 María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández + +! This file was ported from Lean 3 source module ring_theory.mv_polynomial.weighted_homogeneous +! leanprover-community/mathlib commit bcbee715ab85a4f516c814effdf232618c0322af +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Algebra.GradedMonoid +import Mathbin.Data.MvPolynomial.Variables + +/-! +# Weighted homogeneous polynomials + +It is possible to assign weights (in a commutative additive monoid `M`) to the variables of a +multivariate polynomial ring, so that monomials of the ring then have a weighted degree with +respect to the weights of the variables. The weights are represented by a function `w : σ → M`, +where `σ` are the indeterminates. + +A multivariate polynomial `φ` is weighted homogeneous of weighted degree `m : M` if all monomials +occuring in `φ` have the same weighted degree `m`. + +## Main definitions/lemmas + +* `weighted_total_degree' w φ` : the weighted total degree of a multivariate polynomial with respect +to the weights `w`, taking values in `with_bot M`. + +* `weighted_total_degree w φ` : When `M` has a `⊥` element, we can define the weighted total degree +of a multivariate polynomial as a function taking values in `M`. + +* `is_weighted_homogeneous w φ m`: a predicate that asserts that `φ` is weighted homogeneous +of weighted degree `m` with respect to the weights `w`. + +* `weighted_homogeneous_submodule R w m`: the submodule of homogeneous polynomials +of weighted degree `m`. + +* `weighted_homogeneous_component w m`: the additive morphism that projects polynomials +onto their summand that is weighted homogeneous of degree `n` with respect to `w`. + +* `sum_weighted_homogeneous_component`: every polynomial is the sum of its weighted homogeneous +components. +-/ + + +noncomputable section + +open Classical BigOperators + +open Set Function Finset Finsupp AddMonoidAlgebra + +variable {R M : Type _} [CommSemiring R] + +namespace MvPolynomial + +variable {σ : Type _} + +section AddCommMonoid + +variable [AddCommMonoid M] + +/-! ### `weighted_degree'` -/ + + +/-- The `weighted degree'` of the finitely supported function `s : σ →₀ ℕ` is the sum + `∑(s i)•(w i)`. -/ +def weightedDegree' (w : σ → M) : (σ →₀ ℕ) →+ M := + (Finsupp.total σ M ℕ w).toAddMonoidHom +#align mv_polynomial.weighted_degree' MvPolynomial.weightedDegree' + +section SemilatticeSup + +variable [SemilatticeSup M] + +/-- The weighted total degree of a multivariate polynomial, taking values in `with_bot M`. -/ +def weightedTotalDegree' (w : σ → M) (p : MvPolynomial σ R) : WithBot M := + p.support.sup fun s => weightedDegree' w s +#align mv_polynomial.weighted_total_degree' MvPolynomial.weightedTotalDegree' + +/-- The `weighted_total_degree'` of a polynomial `p` is `⊥` if and only if `p = 0`. -/ +theorem weightedTotalDegree'_eq_bot_iff (w : σ → M) (p : MvPolynomial σ R) : + weightedTotalDegree' w p = ⊥ ↔ p = 0 := + by + simp only [weighted_total_degree', Finset.sup_eq_bot_iff, mem_support_iff, WithBot.coe_ne_bot, + MvPolynomial.eq_zero_iff] + exact forall_congr' fun _ => Classical.not_not +#align mv_polynomial.weighted_total_degree'_eq_bot_iff MvPolynomial.weightedTotalDegree'_eq_bot_iff + +/-- The `weighted_total_degree'` of the zero polynomial is `⊥`. -/ +theorem weightedTotalDegree'_zero (w : σ → M) : weightedTotalDegree' w (0 : MvPolynomial σ R) = ⊥ := + by simp only [weighted_total_degree', support_zero, Finset.sup_empty] +#align mv_polynomial.weighted_total_degree'_zero MvPolynomial.weightedTotalDegree'_zero + +section OrderBot + +variable [OrderBot M] + +/-- When `M` has a `⊥` element, we can define the weighted total degree of a multivariate + polynomial as a function taking values in `M`. -/ +def weightedTotalDegree (w : σ → M) (p : MvPolynomial σ R) : M := + p.support.sup fun s => weightedDegree' w s +#align mv_polynomial.weighted_total_degree MvPolynomial.weightedTotalDegree + +/-- This lemma relates `weighted_total_degree` and `weighted_total_degree'`. -/ +theorem weightedTotalDegree_coe (w : σ → M) (p : MvPolynomial σ R) (hp : p ≠ 0) : + weightedTotalDegree' w p = ↑(weightedTotalDegree w p) := + by + rw [Ne.def, ← weighted_total_degree'_eq_bot_iff w p, ← Ne.def, WithBot.ne_bot_iff_exists] at hp + obtain ⟨m, hm⟩ := hp + apply le_antisymm + · simp only [weighted_total_degree, weighted_total_degree', Finset.sup_le_iff, WithBot.coe_le_coe] + intro b + exact Finset.le_sup + · simp only [weighted_total_degree] + have hm' : weighted_total_degree' w p ≤ m := le_of_eq hm.symm + rw [← hm] + simpa [weighted_total_degree'] using hm' +#align mv_polynomial.weighted_total_degree_coe MvPolynomial.weightedTotalDegree_coe + +/-- The `weighted_total_degree` of the zero polynomial is `⊥`. -/ +theorem weightedTotalDegree_zero (w : σ → M) : weightedTotalDegree w (0 : MvPolynomial σ R) = ⊥ := + by simp only [weighted_total_degree, support_zero, Finset.sup_empty] +#align mv_polynomial.weighted_total_degree_zero MvPolynomial.weightedTotalDegree_zero + +theorem le_weightedTotalDegree (w : σ → M) {φ : MvPolynomial σ R} {d : σ →₀ ℕ} + (hd : d ∈ φ.support) : weightedDegree' w d ≤ φ.weightedTotalDegree w := + le_sup hd +#align mv_polynomial.le_weighted_total_degree MvPolynomial.le_weightedTotalDegree + +end OrderBot + +end SemilatticeSup + +/-- A multivariate polynomial `φ` is weighted homogeneous of weighted degree `m` if all monomials + occuring in `φ` have weighted degree `m`. -/ +def IsWeightedHomogeneous (w : σ → M) (φ : MvPolynomial σ R) (m : M) : Prop := + ∀ ⦃d⦄, coeff d φ ≠ 0 → weightedDegree' w d = m +#align mv_polynomial.is_weighted_homogeneous MvPolynomial.IsWeightedHomogeneous + +variable (R) + +/-- The submodule of homogeneous `mv_polynomial`s of degree `n`. -/ +def weightedHomogeneousSubmodule (w : σ → M) (m : M) : Submodule R (MvPolynomial σ R) + where + carrier := { x | x.IsWeightedHomogeneous w m } + smul_mem' r a ha c hc := by + rw [coeff_smul] at hc + exact ha (right_ne_zero_of_mul hc) + zero_mem' d hd := False.elim (hd <| coeff_zero _) + add_mem' a b ha hb c hc := by + rw [coeff_add] at hc + obtain h | h : coeff c a ≠ 0 ∨ coeff c b ≠ 0 := + by + contrapose! hc + simp only [hc, add_zero] + · exact ha h + · exact hb h +#align mv_polynomial.weighted_homogeneous_submodule MvPolynomial.weightedHomogeneousSubmodule + +@[simp] +theorem mem_weightedHomogeneousSubmodule (w : σ → M) (m : M) (p : MvPolynomial σ R) : + p ∈ weightedHomogeneousSubmodule R w m ↔ p.IsWeightedHomogeneous w m := + Iff.rfl +#align mv_polynomial.mem_weighted_homogeneous_submodule MvPolynomial.mem_weightedHomogeneousSubmodule + +variable (R) + +/-- The submodule ` weighted_homogeneous_submodule R w m` of homogeneous `mv_polynomial`s of + degree `n` is equal to the `R`-submodule of all `p : (σ →₀ ℕ) →₀ R` such that + `p.support ⊆ {d | weighted_degree' w d = m}`. While equal, the former has a + convenient definitional reduction. -/ +theorem weightedHomogeneousSubmodule_eq_finsupp_supported (w : σ → M) (m : M) : + weightedHomogeneousSubmodule R w m = Finsupp.supported _ R { d | weightedDegree' w d = m } := + by + ext + simp only [mem_supported, Set.subset_def, Finsupp.mem_support_iff, mem_coe] + rfl +#align mv_polynomial.weighted_homogeneous_submodule_eq_finsupp_supported MvPolynomial.weightedHomogeneousSubmodule_eq_finsupp_supported + +variable {R} + +/-- The submodule generated by products `Pm *Pn` of weighted homogeneous polynomials of degrees `m` + and `n` is contained in the submodule of weighted homogeneous polynomials of degree `m + n`. -/ +theorem weightedHomogeneousSubmodule_mul (w : σ → M) (m n : M) : + weightedHomogeneousSubmodule R w m * weightedHomogeneousSubmodule R w n ≤ + weightedHomogeneousSubmodule R w (m + n) := + by + rw [Submodule.mul_le] + intro φ hφ ψ hψ c hc + rw [coeff_mul] at hc + obtain ⟨⟨d, e⟩, hde, H⟩ := Finset.exists_ne_zero_of_sum_ne_zero hc + have aux : coeff d φ ≠ 0 ∧ coeff e ψ ≠ 0 := + by + contrapose! H + by_cases h : coeff d φ = 0 <;> + simp_all only [Ne.def, not_false_iff, MulZeroClass.zero_mul, MulZeroClass.mul_zero] + rw [← finsupp.mem_antidiagonal.mp hde, ← hφ aux.1, ← hψ aux.2, map_add] +#align mv_polynomial.weighted_homogeneous_submodule_mul MvPolynomial.weightedHomogeneousSubmodule_mul + +/-- Monomials are weighted homogeneous. -/ +theorem isWeightedHomogeneous_monomial (w : σ → M) (d : σ →₀ ℕ) (r : R) {m : M} + (hm : weightedDegree' w d = m) : IsWeightedHomogeneous w (monomial d r) m := + by + intro c hc + rw [coeff_monomial] at hc + split_ifs at hc with h + · subst c + exact hm + · contradiction +#align mv_polynomial.is_weighted_homogeneous_monomial MvPolynomial.isWeightedHomogeneous_monomial + +/-- A polynomial of weighted_total_degree `⊥` is weighted_homogeneous of degree `⊥`. -/ +theorem isWeightedHomogeneous_of_total_degree_zero [SemilatticeSup M] [OrderBot M] (w : σ → M) + {p : MvPolynomial σ R} (hp : weightedTotalDegree w p = (⊥ : M)) : + IsWeightedHomogeneous w p (⊥ : M) := by + intro d hd + have h := weighted_total_degree_coe w p (mv_polynomial.ne_zero_iff.mpr ⟨d, hd⟩) + simp only [weighted_total_degree', hp] at h + rw [eq_bot_iff, ← WithBot.coe_le_coe, ← h] + exact Finset.le_sup (mem_support_iff.mpr hd) +#align mv_polynomial.is_weighted_homogeneous_of_total_degree_zero MvPolynomial.isWeightedHomogeneous_of_total_degree_zero + +/-- Constant polynomials are weighted homogeneous of degree 0. -/ +theorem isWeightedHomogeneous_c (w : σ → M) (r : R) : + IsWeightedHomogeneous w (C r : MvPolynomial σ R) 0 := + isWeightedHomogeneous_monomial _ _ _ (map_zero _) +#align mv_polynomial.is_weighted_homogeneous_C MvPolynomial.isWeightedHomogeneous_c + +variable (R) + +/-- 0 is weighted homogeneous of any degree. -/ +theorem isWeightedHomogeneous_zero (w : σ → M) (m : M) : + IsWeightedHomogeneous w (0 : MvPolynomial σ R) m := + (weightedHomogeneousSubmodule R w m).zero_mem +#align mv_polynomial.is_weighted_homogeneous_zero MvPolynomial.isWeightedHomogeneous_zero + +/-- 1 is weighted homogeneous of degree 0. -/ +theorem isWeightedHomogeneous_one (w : σ → M) : IsWeightedHomogeneous w (1 : MvPolynomial σ R) 0 := + isWeightedHomogeneous_c _ _ +#align mv_polynomial.is_weighted_homogeneous_one MvPolynomial.isWeightedHomogeneous_one + +/-- An indeterminate `i : σ` is weighted homogeneous of degree `w i`. -/ +theorem isWeightedHomogeneous_x (w : σ → M) (i : σ) : + IsWeightedHomogeneous w (X i : MvPolynomial σ R) (w i) := + by + apply is_weighted_homogeneous_monomial + simp only [weighted_degree', LinearMap.toAddMonoidHom_coe, total_single, one_nsmul] +#align mv_polynomial.is_weighted_homogeneous_X MvPolynomial.isWeightedHomogeneous_x + +namespace IsWeightedHomogeneous + +variable {R} {φ ψ : MvPolynomial σ R} {m n : M} + +/-- The weighted degree of a weighted homogeneous polynomial controls its support. -/ +theorem coeff_eq_zero {w : σ → M} (hφ : IsWeightedHomogeneous w φ n) (d : σ →₀ ℕ) + (hd : weightedDegree' w d ≠ n) : coeff d φ = 0 := + by + have aux := mt (@hφ d) hd + rwa [Classical.not_not] at aux +#align mv_polynomial.is_weighted_homogeneous.coeff_eq_zero MvPolynomial.IsWeightedHomogeneous.coeff_eq_zero + +/-- The weighted degree of a nonzero weighted homogeneous polynomial is well-defined. -/ +theorem inj_right {w : σ → M} (hφ : φ ≠ 0) (hm : IsWeightedHomogeneous w φ m) + (hn : IsWeightedHomogeneous w φ n) : m = n := + by + obtain ⟨d, hd⟩ : ∃ d, coeff d φ ≠ 0 := exists_coeff_ne_zero hφ + rw [← hm hd, ← hn hd] +#align mv_polynomial.is_weighted_homogeneous.inj_right MvPolynomial.IsWeightedHomogeneous.inj_right + +/-- The sum of two weighted homogeneous polynomials of degree `n` is weighted homogeneous of + weighted degree `n`. -/ +theorem add {w : σ → M} (hφ : IsWeightedHomogeneous w φ n) (hψ : IsWeightedHomogeneous w ψ n) : + IsWeightedHomogeneous w (φ + ψ) n := + (weightedHomogeneousSubmodule R w n).add_mem hφ hψ +#align mv_polynomial.is_weighted_homogeneous.add MvPolynomial.IsWeightedHomogeneous.add + +/-- The sum of weighted homogeneous polynomials of degree `n` is weighted homogeneous of + weighted degree `n`. -/ +theorem sum {ι : Type _} (s : Finset ι) (φ : ι → MvPolynomial σ R) (n : M) {w : σ → M} + (h : ∀ i ∈ s, IsWeightedHomogeneous w (φ i) n) : IsWeightedHomogeneous w (∑ i in s, φ i) n := + (weightedHomogeneousSubmodule R w n).sum_mem h +#align mv_polynomial.is_weighted_homogeneous.sum MvPolynomial.IsWeightedHomogeneous.sum + +/-- The product of weighted homogeneous polynomials of weighted degrees `m` and `n` is weighted + homogeneous of weighted degree `m + n`. -/ +theorem mul {w : σ → M} (hφ : IsWeightedHomogeneous w φ m) (hψ : IsWeightedHomogeneous w ψ n) : + IsWeightedHomogeneous w (φ * ψ) (m + n) := + weightedHomogeneousSubmodule_mul w m n <| Submodule.mul_mem_mul hφ hψ +#align mv_polynomial.is_weighted_homogeneous.mul MvPolynomial.IsWeightedHomogeneous.mul + +/-- A product of weighted homogeneous polynomials is weighted homogeneous, with weighted degree + equal to the sum of the weighted degrees. -/ +theorem prod {ι : Type _} (s : Finset ι) (φ : ι → MvPolynomial σ R) (n : ι → M) {w : σ → M} : + (∀ i ∈ s, IsWeightedHomogeneous w (φ i) (n i)) → + IsWeightedHomogeneous w (∏ i in s, φ i) (∑ i in s, n i) := + by + apply Finset.induction_on s + · intro + simp only [is_weighted_homogeneous_one, Finset.sum_empty, Finset.prod_empty] + · intro i s his IH h + simp only [his, Finset.prod_insert, Finset.sum_insert, not_false_iff] + apply (h i (Finset.mem_insert_self _ _)).mul (IH _) + intro j hjs + exact h j (Finset.mem_insert_of_mem hjs) +#align mv_polynomial.is_weighted_homogeneous.prod MvPolynomial.IsWeightedHomogeneous.prod + +/-- A non zero weighted homogeneous polynomial of weighted degree `n` has weighted total degree + `n`. -/ +theorem weighted_total_degree [SemilatticeSup M] {w : σ → M} (hφ : IsWeightedHomogeneous w φ n) + (h : φ ≠ 0) : weightedTotalDegree' w φ = n := + by + simp only [weighted_total_degree'] + apply le_antisymm + · simp only [Finset.sup_le_iff, mem_support_iff, WithBot.coe_le_coe] + exact fun d hd => le_of_eq (hφ hd) + · obtain ⟨d, hd⟩ : ∃ d, coeff d φ ≠ 0 := exists_coeff_ne_zero h + simp only [← hφ hd, Finsupp.sum] + replace hd := finsupp.mem_support_iff.mpr hd + exact Finset.le_sup hd +#align mv_polynomial.is_weighted_homogeneous.weighted_total_degree MvPolynomial.IsWeightedHomogeneous.weighted_total_degree + +/-- The weighted homogeneous submodules form a graded monoid. -/ +instance WeightedHomogeneousSubmodule.gcomm_monoid {w : σ → M} : + SetLike.GradedMonoid (weightedHomogeneousSubmodule R w) + where + one_mem := isWeightedHomogeneous_one R w + mul_mem i j xi xj := IsWeightedHomogeneous.mul +#align mv_polynomial.is_weighted_homogeneous.weighted_homogeneous_submodule.gcomm_monoid MvPolynomial.IsWeightedHomogeneous.WeightedHomogeneousSubmodule.gcomm_monoid + +end IsWeightedHomogeneous + +variable {R} + +/-- `weighted_homogeneous_component w n φ` is the part of `φ` that is weighted homogeneous of + weighted degree `n`, with respect to the weights `w`. + See `sum_weighted_homogeneous_component` for the statement that `φ` is equal to the sum + of all its weighted homogeneous components. -/ +def weightedHomogeneousComponent (w : σ → M) (n : M) : MvPolynomial σ R →ₗ[R] MvPolynomial σ R := + (Submodule.subtype _).comp <| Finsupp.restrictDom _ _ { d | weightedDegree' w d = n } +#align mv_polynomial.weighted_homogeneous_component MvPolynomial.weightedHomogeneousComponent + +section WeightedHomogeneousComponent + +variable {w : σ → M} (n : M) (φ ψ : MvPolynomial σ R) + +theorem coeff_weightedHomogeneousComponent (d : σ →₀ ℕ) : + coeff d (weightedHomogeneousComponent w n φ) = + if weightedDegree' w d = n then coeff d φ else 0 := + Finsupp.filter_apply (fun d : σ →₀ ℕ => weightedDegree' w d = n) φ d +#align mv_polynomial.coeff_weighted_homogeneous_component MvPolynomial.coeff_weightedHomogeneousComponent + +theorem weightedHomogeneousComponent_apply : + weightedHomogeneousComponent w n φ = + ∑ d in φ.support.filterₓ fun d => weightedDegree' w d = n, monomial d (coeff d φ) := + Finsupp.filter_eq_sum (fun d : σ →₀ ℕ => weightedDegree' w d = n) φ +#align mv_polynomial.weighted_homogeneous_component_apply MvPolynomial.weightedHomogeneousComponent_apply + +/-- The `n` weighted homogeneous component of a polynomial is weighted homogeneous of +weighted degree `n`. -/ +theorem weightedHomogeneousComponent_isWeightedHomogeneous : + (weightedHomogeneousComponent w n φ).IsWeightedHomogeneous w n := + by + intro d hd + contrapose! hd + rw [coeff_weighted_homogeneous_component, if_neg hd] +#align mv_polynomial.weighted_homogeneous_component_is_weighted_homogeneous MvPolynomial.weightedHomogeneousComponent_isWeightedHomogeneous + +@[simp] +theorem weightedHomogeneousComponent_c_mul (n : M) (r : R) : + weightedHomogeneousComponent w n (C r * φ) = C r * weightedHomogeneousComponent w n φ := by + simp only [C_mul', LinearMap.map_smul] +#align mv_polynomial.weighted_homogeneous_component_C_mul MvPolynomial.weightedHomogeneousComponent_c_mul + +theorem weightedHomogeneousComponent_eq_zero' + (h : ∀ d : σ →₀ ℕ, d ∈ φ.support → weightedDegree' w d ≠ n) : + weightedHomogeneousComponent w n φ = 0 := + by + rw [weighted_homogeneous_component_apply, sum_eq_zero] + intro d hd; rw [mem_filter] at hd + exfalso; exact h _ hd.1 hd.2 +#align mv_polynomial.weighted_homogeneous_component_eq_zero' MvPolynomial.weightedHomogeneousComponent_eq_zero' + +theorem weightedHomogeneousComponent_eq_zero [SemilatticeSup M] [OrderBot M] + (h : weightedTotalDegree w φ < n) : weightedHomogeneousComponent w n φ = 0 := + by + rw [weighted_homogeneous_component_apply, sum_eq_zero] + intro d hd; rw [mem_filter] at hd + exfalso + apply lt_irrefl n + nth_rw 1 [← hd.2] + exact lt_of_le_of_lt (le_weighted_total_degree w hd.1) h +#align mv_polynomial.weighted_homogeneous_component_eq_zero MvPolynomial.weightedHomogeneousComponent_eq_zero + +theorem weightedHomogeneousComponent_finsupp : + (Function.support fun m => weightedHomogeneousComponent w m φ).Finite := + by + suffices + (Function.support fun m => weighted_homogeneous_component w m φ) ⊆ + (fun d => weighted_degree' w d) '' φ.support + by + exact finite.subset ((fun d : σ →₀ ℕ => (weighted_degree' w) d) '' ↑(support φ)).toFinite this + intro m hm + by_contra hm' + apply hm + simp only [mem_support, Ne.def] at hm + simp only [Set.mem_image, not_exists, not_and] at hm' + exact weighted_homogeneous_component_eq_zero' m φ hm' +#align mv_polynomial.weighted_homogeneous_component_finsupp MvPolynomial.weightedHomogeneousComponent_finsupp + +variable (w) + +/-- Every polynomial is the sum of its weighted homogeneous components. -/ +theorem sum_weightedHomogeneousComponent : + (finsum fun m => weightedHomogeneousComponent w m φ) = φ := + by + rw [finsum_eq_sum _ (weighted_homogeneous_component_finsupp φ)] + ext1 d + simp only [coeff_sum, coeff_weighted_homogeneous_component] + rw [Finset.sum_eq_single (weighted_degree' w d)] + · rw [if_pos rfl] + · intro m hm hm' + rw [if_neg hm'.symm] + · intro hm + rw [if_pos rfl] + simp only [finite.mem_to_finset, mem_support, Ne.def, Classical.not_not] at hm + have := coeff_weighted_homogeneous_component _ φ d + rw [hm, if_pos rfl, coeff_zero] at this + exact this.symm +#align mv_polynomial.sum_weighted_homogeneous_component MvPolynomial.sum_weightedHomogeneousComponent + +variable {w} + +/-- The weighted homogeneous components of a weighted homogeneous polynomial. -/ +theorem weightedHomogeneousComponent_weighted_homogeneous_polynomial (m n : M) + (p : MvPolynomial σ R) (h : p ∈ weightedHomogeneousSubmodule R w n) : + weightedHomogeneousComponent w m p = if m = n then p else 0 := + by + simp only [mem_weighted_homogeneous_submodule] at h + ext x + rw [coeff_weighted_homogeneous_component] + by_cases zero_coeff : coeff x p = 0 + · split_ifs + all_goals simp only [zero_coeff, coeff_zero] + · rw [h zero_coeff] + simp only [show n = m ↔ m = n from eq_comm] + split_ifs with h1 + · rfl + · simp only [coeff_zero] +#align mv_polynomial.weighted_homogeneous_component_weighted_homogeneous_polynomial MvPolynomial.weightedHomogeneousComponent_weighted_homogeneous_polynomial + +end WeightedHomogeneousComponent + +end AddCommMonoid + +section CanonicallyOrderedAddMonoid + +variable [CanonicallyOrderedAddMonoid M] {w : σ → M} (φ : MvPolynomial σ R) + +/-- If `M` is a `canonically_ordered_add_monoid`, then the `weighted_homogeneous_component` + of weighted degree `0` of a polynomial is its constant coefficient. -/ +@[simp] +theorem weightedHomogeneousComponent_zero [NoZeroSMulDivisors ℕ M] (hw : ∀ i : σ, w i ≠ 0) : + weightedHomogeneousComponent w 0 φ = C (coeff 0 φ) := + by + ext1 d + rcases em (d = 0) with (rfl | hd) + · simp only [coeff_weighted_homogeneous_component, if_pos, map_zero, coeff_zero_C] + · rw [coeff_weighted_homogeneous_component, if_neg, coeff_C, if_neg (Ne.symm hd)] + simp only [weighted_degree', LinearMap.toAddMonoidHom_coe, Finsupp.total_apply, Finsupp.sum, + sum_eq_zero_iff, Finsupp.mem_support_iff, Ne.def, smul_eq_zero, not_forall, not_or, + and_self_left, exists_prop] + simp only [Finsupp.ext_iff, Finsupp.coe_zero, Pi.zero_apply, not_forall] at hd + obtain ⟨i, hi⟩ := hd + exact ⟨i, hi, hw i⟩ +#align mv_polynomial.weighted_homogeneous_component_zero MvPolynomial.weightedHomogeneousComponent_zero + +end CanonicallyOrderedAddMonoid + +end MvPolynomial + diff --git a/Mathbin/SetTheory/Zfc/Basic.lean b/Mathbin/SetTheory/Zfc/Basic.lean index 68d0bde2c6..7ba153d34f 100644 --- a/Mathbin/SetTheory/Zfc/Basic.lean +++ b/Mathbin/SetTheory/Zfc/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro ! This file was ported from Lean 3 source module set_theory.zfc.basic -! leanprover-community/mathlib commit ef5f2ce93dd79b7fb184018b6f48132a10c764e7 +! leanprover-community/mathlib commit 98bbc3526516bca903bff09ea10c4206bf079e6b ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -1168,6 +1168,21 @@ theorem mem_diff {x y z : SetCat.{u}} : z ∈ x \ y ↔ z ∈ x ∧ z ∉ y := @mem_sep fun z : SetCat.{u} => z ∉ y #align Set.mem_diff SetCat.mem_diff +/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ +@[simp] +theorem sUnion_pair {x y : SetCat.{u}} : ⋃₀ ({x, y} : SetCat.{u}) = x ∪ y := + by + ext + simp_rw [mem_union, mem_sUnion, mem_pair] + constructor + · rintro ⟨w, rfl | rfl, hw⟩ + · exact Or.inl hw + · exact Or.inr hw + · rintro (hz | hz) + · exact ⟨x, Or.inl rfl, hz⟩ + · exact ⟨y, Or.inr rfl, hz⟩ +#align Set.sUnion_pair SetCat.sUnion_pair + theorem mem_wf : @WellFounded SetCat (· ∈ ·) := wellFounded_lift₂_iff.mpr PSet.mem_wf #align Set.mem_wf SetCat.mem_wf diff --git a/Mathbin/SetTheory/Zfc/Ordinal.lean b/Mathbin/SetTheory/Zfc/Ordinal.lean index 90172b3af6..4fece11d0a 100644 --- a/Mathbin/SetTheory/Zfc/Ordinal.lean +++ b/Mathbin/SetTheory/Zfc/Ordinal.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios ! This file was ported from Lean 3 source module set_theory.zfc.ordinal -! leanprover-community/mathlib commit 7a1cc03e365059994e64dc313bb023427bcfb50d +! leanprover-community/mathlib commit 98bbc3526516bca903bff09ea10c4206bf079e6b ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -78,6 +78,21 @@ theorem IsTransitive.sUnion' (H : ∀ y ∈ x, IsTransitive y) : (⋃₀ x).IsTr exact mem_sUnion_of_mem ((H w hw).mem_trans hz hw') hw #align Set.is_transitive.sUnion' SetCat.IsTransitive.sUnion' +protected theorem IsTransitive.union (hx : x.IsTransitive) (hy : y.IsTransitive) : + (x ∪ y).IsTransitive := by + rw [← sUnion_pair] + apply is_transitive.sUnion' fun z => _ + rw [mem_pair] + rintro (rfl | rfl) + assumption' +#align Set.is_transitive.union SetCat.IsTransitive.union + +protected theorem IsTransitive.powerset (h : x.IsTransitive) : (powerset x).IsTransitive := + fun y hy z hz => by + rw [mem_powerset] at hy⊢ + exact h.subset_of_mem (hy hz) +#align Set.is_transitive.powerset SetCat.IsTransitive.powerset + /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem isTransitive_iff_sUnion_subset : x.IsTransitive ↔ ⋃₀ x ⊆ x := ⟨fun h y hy => by diff --git a/Mathbin/Topology/Category/Locale.lean b/Mathbin/Topology/Category/Locale.lean index f219779aaf..1bb7924a3c 100644 --- a/Mathbin/Topology/Category/Locale.lean +++ b/Mathbin/Topology/Category/Locale.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies ! This file was ported from Lean 3 source module topology.category.Locale -! leanprover-community/mathlib commit 5a9fb92b86469f4111a69819ac2ee21d0de3dfd0 +! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.Order.Category.Frame +import Mathbin.Order.Category.Frm /-! # The category of locales @@ -23,7 +23,7 @@ open CategoryTheory Opposite Order TopologicalSpace /-- The category of locales. -/ def Locale := - Frameᵒᵖderiving LargeCategory + Frmᵒᵖderiving LargeCategory #align Locale Locale namespace Locale @@ -36,7 +36,7 @@ instance (X : Locale) : Frame X := /-- Construct a bundled `Locale` from a `frame`. -/ def of (α : Type _) [Frame α] : Locale := - op <| Frame.of α + op <| Frm.of α #align Locale.of Locale.of @[simp] diff --git a/Mathbin/Topology/FiberBundle/Basic.lean b/Mathbin/Topology/FiberBundle/Basic.lean index d317ceff03..d740ed8e53 100644 --- a/Mathbin/Topology/FiberBundle/Basic.lean +++ b/Mathbin/Topology/FiberBundle/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn, Heather Macbeth ! This file was ported from Lean 3 source module topology.fiber_bundle.basic -! leanprover-community/mathlib commit 8ef6f08ff8c781c5c07a8b12843710e1a0d8a688 +! leanprover-community/mathlib commit 0187644979f2d3e10a06e916a869c994facd9a87 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -173,14 +173,14 @@ Fiber bundle, topological bundle, structure group -/ -variable {ι : Type _} {B : Type _} {F : Type _} +variable {ι B F X : Type _} [TopologicalSpace X] open TopologicalSpace Filter Set Bundle open Topology Classical Bundle attribute [mfld_simps] - total_space.proj total_space_mk coe_fst coe_snd coe_snd_map_apply coe_snd_map_smul total_space.mk_cast + total_space_mk coe_fst coe_snd coe_snd_map_apply coe_snd_map_smul total_space.mk_cast /-! ### General definition of fiber bundles -/ @@ -233,9 +233,9 @@ variable (F) {E} [FiberBundle F E] /- warning: fiber_bundle.map_proj_nhds -> FiberBundle.map_proj_nhds is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} (F : Type.{u2}) [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] {E : B -> Type.{u3}} [_inst_3 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_5 : FiberBundle.{u1, u2, u3} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] (x : Bundle.TotalSpace.{u1, u3} B E), Eq.{succ u1} (Filter.{u1} B) (Filter.map.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B (Bundle.TotalSpace.proj.{u1, u3} B E) (nhds.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E) _inst_3 x)) (nhds.{u1} B _inst_1 (Bundle.TotalSpace.proj.{u1, u3} B E x)) + forall {B : Type.{u1}} (F : Type.{u2}) [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] {E : B -> Type.{u3}} [_inst_4 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_6 : FiberBundle.{u1, u2, u3} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] (x : Bundle.TotalSpace.{u1, u3} B E), Eq.{succ u1} (Filter.{u1} B) (Filter.map.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B (Bundle.TotalSpace.proj.{u1, u3} B E) (nhds.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E) _inst_4 x)) (nhds.{u1} B _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E x)) but is expected to have type - forall {B : Type.{u2}} (F : Type.{u3}) [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] {E : B -> Type.{u1}} [_inst_3 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_5 : FiberBundle.{u2, u3, u1} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] (x : Bundle.TotalSpace.{u2, u1} B E), Eq.{succ u2} (Filter.{u2} B) (Filter.map.{max u2 u1, u2} (Bundle.TotalSpace.{u2, u1} B E) B (Bundle.TotalSpace.proj.{u2, u1} B E) (nhds.{max u2 u1} (Bundle.TotalSpace.{u2, u1} B E) _inst_3 x)) (nhds.{u2} B _inst_1 (Bundle.TotalSpace.proj.{u2, u1} B E x)) + forall {B : Type.{u2}} (F : Type.{u3}) [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] {E : B -> Type.{u1}} [_inst_4 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_6 : FiberBundle.{u2, u3, u1} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] (x : Bundle.TotalSpace.{u2, u1} B E), Eq.{succ u2} (Filter.{u2} B) (Filter.map.{max u2 u1, u2} (Bundle.TotalSpace.{u2, u1} B E) B (Bundle.TotalSpace.proj.{u2, u1} B E) (nhds.{max u2 u1} (Bundle.TotalSpace.{u2, u1} B E) _inst_4 x)) (nhds.{u2} B _inst_2 (Bundle.TotalSpace.proj.{u2, u1} B E x)) Case conversion may be inaccurate. Consider using '#align fiber_bundle.map_proj_nhds FiberBundle.map_proj_nhdsₓ'. -/ theorem map_proj_nhds (x : TotalSpace E) : map (π E) (𝓝 x) = 𝓝 x.proj := (trivializationAt F E x.proj).map_proj_nhds <| @@ -246,9 +246,9 @@ variable (E) /- warning: fiber_bundle.continuous_proj -> FiberBundle.continuous_proj is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} (F : Type.{u2}) [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_3 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_5 : FiberBundle.{u1, u2, u3} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)], Continuous.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B _inst_3 _inst_1 (Bundle.TotalSpace.proj.{u1, u3} B E) + forall {B : Type.{u1}} (F : Type.{u2}) [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_4 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_6 : FiberBundle.{u1, u2, u3} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)], Continuous.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B _inst_4 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) but is expected to have type - forall {B : Type.{u2}} (F : Type.{u3}) [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (E : B -> Type.{u1}) [_inst_3 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_5 : FiberBundle.{u2, u3, u1} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)], Continuous.{max u2 u1, u2} (Bundle.TotalSpace.{u2, u1} B E) B _inst_3 _inst_1 (Bundle.TotalSpace.proj.{u2, u1} B E) + forall {B : Type.{u2}} (F : Type.{u3}) [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (E : B -> Type.{u1}) [_inst_4 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_6 : FiberBundle.{u2, u3, u1} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)], Continuous.{max u2 u1, u2} (Bundle.TotalSpace.{u2, u1} B E) B _inst_4 _inst_2 (Bundle.TotalSpace.proj.{u2, u1} B E) Case conversion may be inaccurate. Consider using '#align fiber_bundle.continuous_proj FiberBundle.continuous_projₓ'. -/ /-- The projection from a fiber bundle to its base is continuous. -/ @[continuity] @@ -258,9 +258,9 @@ theorem continuous_proj : Continuous (π E) := /- warning: fiber_bundle.is_open_map_proj -> FiberBundle.isOpenMap_proj is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} (F : Type.{u2}) [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_3 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_5 : FiberBundle.{u1, u2, u3} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)], IsOpenMap.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B _inst_3 _inst_1 (Bundle.TotalSpace.proj.{u1, u3} B E) + forall {B : Type.{u1}} (F : Type.{u2}) [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_4 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_6 : FiberBundle.{u1, u2, u3} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)], IsOpenMap.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B _inst_4 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) but is expected to have type - forall {B : Type.{u2}} (F : Type.{u3}) [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (E : B -> Type.{u1}) [_inst_3 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_5 : FiberBundle.{u2, u3, u1} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)], IsOpenMap.{max u2 u1, u2} (Bundle.TotalSpace.{u2, u1} B E) B _inst_3 _inst_1 (Bundle.TotalSpace.proj.{u2, u1} B E) + forall {B : Type.{u2}} (F : Type.{u3}) [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (E : B -> Type.{u1}) [_inst_4 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_6 : FiberBundle.{u2, u3, u1} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)], IsOpenMap.{max u2 u1, u2} (Bundle.TotalSpace.{u2, u1} B E) B _inst_4 _inst_2 (Bundle.TotalSpace.proj.{u2, u1} B E) Case conversion may be inaccurate. Consider using '#align fiber_bundle.is_open_map_proj FiberBundle.isOpenMap_projₓ'. -/ /-- The projection from a fiber bundle to its base is an open map. -/ theorem isOpenMap_proj : IsOpenMap (π E) := @@ -269,9 +269,9 @@ theorem isOpenMap_proj : IsOpenMap (π E) := /- warning: fiber_bundle.surjective_proj -> FiberBundle.surjective_proj is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} (F : Type.{u2}) [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_3 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_5 : FiberBundle.{u1, u2, u3} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] [_inst_6 : Nonempty.{succ u2} F], Function.Surjective.{max (succ u1) (succ u3), succ u1} (Bundle.TotalSpace.{u1, u3} B E) B (Bundle.TotalSpace.proj.{u1, u3} B E) + forall {B : Type.{u1}} (F : Type.{u2}) [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_4 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_6 : FiberBundle.{u1, u2, u3} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] [_inst_7 : Nonempty.{succ u2} F], Function.Surjective.{max (succ u1) (succ u3), succ u1} (Bundle.TotalSpace.{u1, u3} B E) B (Bundle.TotalSpace.proj.{u1, u3} B E) but is expected to have type - forall {B : Type.{u2}} (F : Type.{u3}) [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (E : B -> Type.{u1}) [_inst_3 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_5 : FiberBundle.{u2, u3, u1} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] [_inst_6 : Nonempty.{succ u3} F], Function.Surjective.{max (succ u2) (succ u1), succ u2} (Bundle.TotalSpace.{u2, u1} B E) B (Bundle.TotalSpace.proj.{u2, u1} B E) + forall {B : Type.{u2}} (F : Type.{u3}) [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (E : B -> Type.{u1}) [_inst_4 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_6 : FiberBundle.{u2, u3, u1} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] [_inst_7 : Nonempty.{succ u3} F], Function.Surjective.{max (succ u2) (succ u1), succ u2} (Bundle.TotalSpace.{u2, u1} B E) B (Bundle.TotalSpace.proj.{u2, u1} B E) Case conversion may be inaccurate. Consider using '#align fiber_bundle.surjective_proj FiberBundle.surjective_projₓ'. -/ /-- The projection from a fiber bundle with a nonempty fiber to its base is a surjective map. -/ @@ -283,9 +283,9 @@ theorem surjective_proj [Nonempty F] : Function.Surjective (π E) := fun b => /- warning: fiber_bundle.quotient_map_proj -> FiberBundle.quotientMap_proj is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} (F : Type.{u2}) [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_3 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_5 : FiberBundle.{u1, u2, u3} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] [_inst_6 : Nonempty.{succ u2} F], QuotientMap.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B _inst_3 _inst_1 (Bundle.TotalSpace.proj.{u1, u3} B E) + forall {B : Type.{u1}} (F : Type.{u2}) [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_4 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_6 : FiberBundle.{u1, u2, u3} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] [_inst_7 : Nonempty.{succ u2} F], QuotientMap.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B _inst_4 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) but is expected to have type - forall {B : Type.{u2}} (F : Type.{u3}) [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (E : B -> Type.{u1}) [_inst_3 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_5 : FiberBundle.{u2, u3, u1} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] [_inst_6 : Nonempty.{succ u3} F], QuotientMap.{max u2 u1, u2} (Bundle.TotalSpace.{u2, u1} B E) B _inst_3 _inst_1 (Bundle.TotalSpace.proj.{u2, u1} B E) + forall {B : Type.{u2}} (F : Type.{u3}) [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (E : B -> Type.{u1}) [_inst_4 : TopologicalSpace.{max u1 u2} (Bundle.TotalSpace.{u2, u1} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_6 : FiberBundle.{u2, u3, u1} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] [_inst_7 : Nonempty.{succ u3} F], QuotientMap.{max u2 u1, u2} (Bundle.TotalSpace.{u2, u1} B E) B _inst_4 _inst_2 (Bundle.TotalSpace.proj.{u2, u1} B E) Case conversion may be inaccurate. Consider using '#align fiber_bundle.quotient_map_proj FiberBundle.quotientMap_projₓ'. -/ /-- The projection from a fiber bundle with a nonempty fiber to its base is a quotient map. -/ @@ -295,23 +295,73 @@ theorem quotientMap_proj [Nonempty F] : QuotientMap (π E) := /- warning: fiber_bundle.continuous_total_space_mk -> FiberBundle.continuous_totalSpaceMk is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} (F : Type.{u2}) [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_3 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_5 : FiberBundle.{u1, u2, u3} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] (x : B), Continuous.{u3, max u1 u3} (E x) (Bundle.TotalSpace.{u1, u3} B E) (_inst_4 x) _inst_3 (Bundle.totalSpaceMk.{u1, u3} B E x) + forall {B : Type.{u1}} (F : Type.{u2}) [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_4 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_6 : FiberBundle.{u1, u2, u3} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] (x : B), Continuous.{u3, max u1 u3} (E x) (Bundle.TotalSpace.{u1, u3} B E) (_inst_5 x) _inst_4 (Bundle.totalSpaceMk.{u1, u3} B E x) but is expected to have type - forall {B : Type.{u1}} (F : Type.{u3}) [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u3} F] (E : B -> Type.{u2}) [_inst_3 : TopologicalSpace.{max u2 u1} (Bundle.TotalSpace.{u1, u2} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u2} (E b)] [_inst_5 : FiberBundle.{u1, u3, u2} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] (x : B), Continuous.{u2, max u1 u2} (E x) (Bundle.TotalSpace.{u1, u2} B E) (_inst_4 x) _inst_3 (Bundle.totalSpaceMk.{u1, u2} B E x) + forall {B : Type.{u1}} (F : Type.{u3}) [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u3} F] (E : B -> Type.{u2}) [_inst_4 : TopologicalSpace.{max u2 u1} (Bundle.TotalSpace.{u1, u2} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u2} (E b)] [_inst_6 : FiberBundle.{u1, u3, u2} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] (x : B), Continuous.{u2, max u1 u2} (E x) (Bundle.TotalSpace.{u1, u2} B E) (_inst_5 x) _inst_4 (Bundle.totalSpaceMk.{u1, u2} B E x) Case conversion may be inaccurate. Consider using '#align fiber_bundle.continuous_total_space_mk FiberBundle.continuous_totalSpaceMkₓ'. -/ theorem continuous_totalSpaceMk (x : B) : Continuous (@totalSpaceMk B E x) := (totalSpaceMk_inducing F E x).Continuous #align fiber_bundle.continuous_total_space_mk FiberBundle.continuous_totalSpaceMk +variable {E F} + +@[simp, mfld_simps] +theorem mem_trivializationAt_proj_source {x : TotalSpace E} : + x ∈ (trivializationAt F E x.proj).source := + (Trivialization.mem_source _).mpr <| mem_baseSet_trivializationAt F E x.proj +#align fiber_bundle.mem_trivialization_at_proj_source FiberBundle.mem_trivializationAt_proj_source + +@[simp, mfld_simps] +theorem trivializationAt_proj_fst {x : TotalSpace E} : + ((trivializationAt F E x.proj) x).1 = x.proj := + Trivialization.coe_fst' _ <| mem_baseSet_trivializationAt F E x.proj +#align fiber_bundle.trivialization_at_proj_fst FiberBundle.trivializationAt_proj_fst + +variable (F) + +open Trivialization + +/-- Characterization of continuous functions (at a point, within a set) into a fiber bundle. -/ +theorem continuousWithinAt_totalSpace (f : X → TotalSpace E) {s : Set X} {x₀ : X} : + ContinuousWithinAt f s x₀ ↔ + ContinuousWithinAt (fun x => (f x).proj) s x₀ ∧ + ContinuousWithinAt (fun x => ((trivializationAt F E (f x₀).proj) (f x)).2) s x₀ := + by + refine' (and_iff_right_iff_imp.2 fun hf => _).symm.trans (and_congr_right fun hf => _) + · refine' (continuous_proj F E).ContinuousWithinAt.comp hf (maps_to_image f s) + have h1 : (fun x => (f x).proj) ⁻¹' (trivialization_at F E (f x₀).proj).baseSet ∈ 𝓝[s] x₀ := + hf.preimage_mem_nhds_within ((open_base_set _).mem_nhds (mem_base_set_trivialization_at F E _)) + have h2 : ContinuousWithinAt (fun x => (trivialization_at F E (f x₀).proj (f x)).1) s x₀ := + by + refine' + hf.congr_of_eventually_eq (eventually_of_mem h1 fun x hx => _) trivialization_at_proj_fst + rw [coe_fst'] + exact hx + rw [(trivialization_at F E (f x₀).proj).continuousWithinAt_iff_continuousWithinAt_comp_left] + · simp_rw [continuousWithinAt_prod_iff, Function.comp, Trivialization.coe_coe, h2, true_and_iff] + · apply mem_trivialization_at_proj_source + · rwa [source_eq, preimage_preimage] +#align fiber_bundle.continuous_within_at_total_space FiberBundle.continuousWithinAt_totalSpace + +/-- Characterization of continuous functions (at a point) into a fiber bundle. -/ +theorem continuousAt_totalSpace (f : X → TotalSpace E) {x₀ : X} : + ContinuousAt f x₀ ↔ + ContinuousAt (fun x => (f x).proj) x₀ ∧ + ContinuousAt (fun x => ((trivializationAt F E (f x₀).proj) (f x)).2) x₀ := + by + simp_rw [← continuousWithinAt_univ] + exact continuous_within_at_total_space F f +#align fiber_bundle.continuous_at_total_space FiberBundle.continuousAt_totalSpace + end FiberBundle variable (F E) /- warning: fiber_bundle.exists_trivialization_Icc_subset -> FiberBundle.exists_trivialization_Icc_subset is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} (F : Type.{u2}) [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_3 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_5 : ConditionallyCompleteLinearOrder.{u1} B] [_inst_6 : OrderTopology.{u1} B _inst_1 (PartialOrder.toPreorder.{u1} B (SemilatticeInf.toPartialOrder.{u1} B (Lattice.toSemilatticeInf.{u1} B (ConditionallyCompleteLattice.toLattice.{u1} B (ConditionallyCompleteLinearOrder.toConditionallyCompleteLattice.{u1} B _inst_5)))))] [_inst_7 : FiberBundle.{u1, u2, u3} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] (a : B) (b : B), Exists.{max (succ u1) (succ u2) (succ (max u1 u3))} (Trivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E)) (fun (e : Trivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E)) => HasSubset.Subset.{u1} (Set.{u1} B) (Set.hasSubset.{u1} B) (Set.Icc.{u1} B (PartialOrder.toPreorder.{u1} B (SemilatticeInf.toPartialOrder.{u1} B (Lattice.toSemilatticeInf.{u1} B (ConditionallyCompleteLattice.toLattice.{u1} B (ConditionallyCompleteLinearOrder.toConditionallyCompleteLattice.{u1} B _inst_5))))) a b) (Trivialization.baseSet.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) e)) + forall {B : Type.{u1}} (F : Type.{u2}) [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (E : B -> Type.{u3}) [_inst_4 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u3} (E b)] [_inst_6 : ConditionallyCompleteLinearOrder.{u1} B] [_inst_7 : OrderTopology.{u1} B _inst_2 (PartialOrder.toPreorder.{u1} B (SemilatticeInf.toPartialOrder.{u1} B (Lattice.toSemilatticeInf.{u1} B (ConditionallyCompleteLattice.toLattice.{u1} B (ConditionallyCompleteLinearOrder.toConditionallyCompleteLattice.{u1} B _inst_6)))))] [_inst_8 : FiberBundle.{u1, u2, u3} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] (a : B) (b : B), Exists.{max (succ u1) (succ u2) (succ (max u1 u3))} (Trivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 _inst_4 (Bundle.TotalSpace.proj.{u1, u3} B E)) (fun (e : Trivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 _inst_4 (Bundle.TotalSpace.proj.{u1, u3} B E)) => HasSubset.Subset.{u1} (Set.{u1} B) (Set.hasSubset.{u1} B) (Set.Icc.{u1} B (PartialOrder.toPreorder.{u1} B (SemilatticeInf.toPartialOrder.{u1} B (Lattice.toSemilatticeInf.{u1} B (ConditionallyCompleteLattice.toLattice.{u1} B (ConditionallyCompleteLinearOrder.toConditionallyCompleteLattice.{u1} B _inst_6))))) a b) (Trivialization.baseSet.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 _inst_4 (Bundle.TotalSpace.proj.{u1, u3} B E) e)) but is expected to have type - forall {B : Type.{u3}} (F : Type.{u2}) [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (E : B -> Type.{u1}) [_inst_3 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u3, u1} B E)] [_inst_4 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_5 : ConditionallyCompleteLinearOrder.{u3} B] [_inst_6 : OrderTopology.{u3} B _inst_1 (PartialOrder.toPreorder.{u3} B (SemilatticeInf.toPartialOrder.{u3} B (Lattice.toSemilatticeInf.{u3} B (ConditionallyCompleteLattice.toLattice.{u3} B (ConditionallyCompleteLinearOrder.toConditionallyCompleteLattice.{u3} B _inst_5)))))] [_inst_7 : FiberBundle.{u3, u2, u1} B F _inst_1 _inst_2 E _inst_3 (fun (b : B) => _inst_4 b)] (a : B) (b : B), Exists.{max (max (succ u3) (succ u2)) (succ u1)} (Trivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E)) (fun (e : Trivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E)) => HasSubset.Subset.{u3} (Set.{u3} B) (Set.instHasSubsetSet.{u3} B) (Set.Icc.{u3} B (PartialOrder.toPreorder.{u3} B (SemilatticeInf.toPartialOrder.{u3} B (Lattice.toSemilatticeInf.{u3} B (ConditionallyCompleteLattice.toLattice.{u3} B (ConditionallyCompleteLinearOrder.toConditionallyCompleteLattice.{u3} B _inst_5))))) a b) (Trivialization.baseSet.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E) e)) + forall {B : Type.{u3}} (F : Type.{u2}) [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (E : B -> Type.{u1}) [_inst_4 : TopologicalSpace.{max u1 u3} (Bundle.TotalSpace.{u3, u1} B E)] [_inst_5 : forall (b : B), TopologicalSpace.{u1} (E b)] [_inst_6 : ConditionallyCompleteLinearOrder.{u3} B] [_inst_7 : OrderTopology.{u3} B _inst_2 (PartialOrder.toPreorder.{u3} B (SemilatticeInf.toPartialOrder.{u3} B (Lattice.toSemilatticeInf.{u3} B (ConditionallyCompleteLattice.toLattice.{u3} B (ConditionallyCompleteLinearOrder.toConditionallyCompleteLattice.{u3} B _inst_6)))))] [_inst_8 : FiberBundle.{u3, u2, u1} B F _inst_2 _inst_3 E _inst_4 (fun (b : B) => _inst_5 b)] (a : B) (b : B), Exists.{max (max (succ u3) (succ u2)) (succ u1)} (Trivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 _inst_4 (Bundle.TotalSpace.proj.{u3, u1} B E)) (fun (e : Trivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 _inst_4 (Bundle.TotalSpace.proj.{u3, u1} B E)) => HasSubset.Subset.{u3} (Set.{u3} B) (Set.instHasSubsetSet.{u3} B) (Set.Icc.{u3} B (PartialOrder.toPreorder.{u3} B (SemilatticeInf.toPartialOrder.{u3} B (Lattice.toSemilatticeInf.{u3} B (ConditionallyCompleteLattice.toLattice.{u3} B (ConditionallyCompleteLinearOrder.toConditionallyCompleteLattice.{u3} B _inst_6))))) a b) (Trivialization.baseSet.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 _inst_4 (Bundle.TotalSpace.proj.{u3, u1} B E) e)) Case conversion may be inaccurate. Consider using '#align fiber_bundle.exists_trivialization_Icc_subset FiberBundle.exists_trivialization_Icc_subsetₓ'. -/ /-- If `E` is a fiber bundle over a conditionally complete linear order, then it is trivial over any closed interval. -/ @@ -482,9 +532,9 @@ def proj : Z.TotalSpace → B := /- warning: fiber_bundle_core.triv_change -> FiberBundleCore.trivChange is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F], (FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) -> ι -> ι -> (LocalHomeomorph.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F], (FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) -> ι -> ι -> (LocalHomeomorph.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3)) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F], (FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) -> ι -> ι -> (LocalHomeomorph.{max u3 u2, max u3 u2} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (instTopologicalSpaceProd.{u2, u3} B F _inst_1 _inst_2) (instTopologicalSpaceProd.{u2, u3} B F _inst_1 _inst_2)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F], (FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) -> ι -> ι -> (LocalHomeomorph.{max u3 u2, max u3 u2} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (instTopologicalSpaceProd.{u2, u3} B F _inst_2 _inst_3) (instTopologicalSpaceProd.{u2, u3} B F _inst_2 _inst_3)) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.triv_change FiberBundleCore.trivChangeₓ'. -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @@ -519,9 +569,9 @@ def trivChange (i j : ι) : LocalHomeomorph (B × F) (B × F) /- warning: fiber_bundle_core.mem_triv_change_source -> FiberBundleCore.mem_trivChange_source is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι) (j : ι) (p : Prod.{u2, u3} B F), Iff (Membership.Mem.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Set.{max u2 u3} (Prod.{u2, u3} B F)) (Set.hasMem.{max u2 u3} (Prod.{u2, u3} B F)) p (LocalEquiv.source.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (FiberBundleCore.trivChange.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i j)))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Prod.fst.{u2, u3} B F p) (Inter.inter.{u2} (Set.{u2} B) (Set.hasInter.{u2} B) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_1 F _inst_2 Z i) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_1 F _inst_2 Z j))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι) (j : ι) (p : Prod.{u2, u3} B F), Iff (Membership.Mem.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Set.{max u2 u3} (Prod.{u2, u3} B F)) (Set.hasMem.{max u2 u3} (Prod.{u2, u3} B F)) p (LocalEquiv.source.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (FiberBundleCore.trivChange.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i j)))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Prod.fst.{u2, u3} B F p) (Inter.inter.{u2} (Set.{u2} B) (Set.hasInter.{u2} B) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_2 F _inst_3 Z i) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_2 F _inst_3 Z j))) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι) (j : ι) (p : Prod.{u3, u2} B F), Iff (Membership.mem.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instMembershipSet.{max u3 u2} (Prod.{u3, u2} B F)) p (LocalEquiv.source.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (FiberBundleCore.trivChange.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i j)))) (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) (Prod.fst.{u3, u2} B F p) (Inter.inter.{u3} (Set.{u3} B) (Set.instInterSet.{u3} B) (FiberBundleCore.baseSet.{u1, u3, u2} ι B _inst_1 F _inst_2 Z i) (FiberBundleCore.baseSet.{u1, u3, u2} ι B _inst_1 F _inst_2 Z j))) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι) (j : ι) (p : Prod.{u3, u2} B F), Iff (Membership.mem.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instMembershipSet.{max u3 u2} (Prod.{u3, u2} B F)) p (LocalEquiv.source.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (FiberBundleCore.trivChange.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i j)))) (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) (Prod.fst.{u3, u2} B F p) (Inter.inter.{u3} (Set.{u3} B) (Set.instInterSet.{u3} B) (FiberBundleCore.baseSet.{u1, u3, u2} ι B _inst_2 F _inst_3 Z i) (FiberBundleCore.baseSet.{u1, u3, u2} ι B _inst_2 F _inst_3 Z j))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.mem_triv_change_source FiberBundleCore.mem_trivChange_sourceₓ'. -/ @[simp, mfld_simps] theorem mem_trivChange_source (i j : ι) (p : B × F) : @@ -571,9 +621,9 @@ variable (i : ι) /- warning: fiber_bundle_core.mem_local_triv_as_local_equiv_source -> FiberBundleCore.mem_localTrivAsLocalEquiv_source is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z), Iff (Membership.Mem.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Set.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (Set.hasMem.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) p (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_1 F _inst_2 Z i)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z), Iff (Membership.Mem.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Set.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (Set.hasMem.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) p (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_2 F _inst_3 Z i)) but is expected to have type - forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_1 F _inst_2) (i : ι) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z), Iff (Membership.mem.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (Set.{max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) (Set.instMembershipSet.{max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) p (LocalEquiv.source.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u1} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u3, u2, u1} ι B F _inst_1 _inst_2 Z i))) (Membership.mem.{u2, u2} B (Set.{u2} B) (Set.instMembershipSet.{u2} B) (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p) (FiberBundleCore.baseSet.{u3, u2, u1} ι B _inst_1 F _inst_2 Z i)) + forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_2 F _inst_3) (i : ι) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z), Iff (Membership.mem.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (Set.{max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) (Set.instMembershipSet.{max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) p (LocalEquiv.source.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u1} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u3, u2, u1} ι B F _inst_2 _inst_3 Z i))) (Membership.mem.{u2, u2} B (Set.{u2} B) (Set.instMembershipSet.{u2} B) (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p) (FiberBundleCore.baseSet.{u3, u2, u1} ι B _inst_2 F _inst_3 Z i)) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.mem_local_triv_as_local_equiv_source FiberBundleCore.mem_localTrivAsLocalEquiv_sourceₓ'. -/ theorem mem_localTrivAsLocalEquiv_source (p : Z.TotalSpace) : p ∈ (Z.localTrivAsLocalEquiv i).source ↔ p.1 ∈ Z.baseSet i := @@ -582,9 +632,9 @@ theorem mem_localTrivAsLocalEquiv_source (p : Z.TotalSpace) : /- warning: fiber_bundle_core.mem_local_triv_as_local_equiv_target -> FiberBundleCore.mem_localTrivAsLocalEquiv_target is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι) (p : Prod.{u2, u3} B F), Iff (Membership.Mem.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Set.{max u2 u3} (Prod.{u2, u3} B F)) (Set.hasMem.{max u2 u3} (Prod.{u2, u3} B F)) p (LocalEquiv.target.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Prod.fst.{u2, u3} B F p) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_1 F _inst_2 Z i)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι) (p : Prod.{u2, u3} B F), Iff (Membership.Mem.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Set.{max u2 u3} (Prod.{u2, u3} B F)) (Set.hasMem.{max u2 u3} (Prod.{u2, u3} B F)) p (LocalEquiv.target.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Prod.fst.{u2, u3} B F p) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_2 F _inst_3 Z i)) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι) (p : Prod.{u3, u2} B F), Iff (Membership.mem.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instMembershipSet.{max u3 u2} (Prod.{u3, u2} B F)) p (LocalEquiv.target.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i))) (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) (Prod.fst.{u3, u2} B F p) (FiberBundleCore.baseSet.{u1, u3, u2} ι B _inst_1 F _inst_2 Z i)) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι) (p : Prod.{u3, u2} B F), Iff (Membership.mem.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instMembershipSet.{max u3 u2} (Prod.{u3, u2} B F)) p (LocalEquiv.target.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i))) (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) (Prod.fst.{u3, u2} B F p) (FiberBundleCore.baseSet.{u1, u3, u2} ι B _inst_2 F _inst_3 Z i)) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.mem_local_triv_as_local_equiv_target FiberBundleCore.mem_localTrivAsLocalEquiv_targetₓ'. -/ theorem mem_localTrivAsLocalEquiv_target (p : B × F) : p ∈ (Z.localTrivAsLocalEquiv i).target ↔ p.1 ∈ Z.baseSet i := @@ -595,9 +645,9 @@ theorem mem_localTrivAsLocalEquiv_target (p : B × F) : /- warning: fiber_bundle_core.local_triv_as_local_equiv_apply -> FiberBundleCore.localTrivAsLocalEquiv_apply is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z), Eq.{max (succ u2) (succ u3)} (Prod.{u2, u3} B F) (coeFn.{succ (max u2 u3), succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F)) (fun (_x : LocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F)) => (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) -> (Prod.{u2, u3} B F)) (LocalEquiv.hasCoeToFun.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F)) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i) p) (Prod.mk.{u2, u3} B F (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p) (FiberBundleCore.coordChange.{u1, u2, u3} ι B _inst_1 F _inst_2 Z (FiberBundleCore.indexAt.{u1, u2, u3} ι B _inst_1 F _inst_2 Z (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p)) i (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p) (Sigma.snd.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z), Eq.{max (succ u2) (succ u3)} (Prod.{u2, u3} B F) (coeFn.{succ (max u2 u3), succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F)) (fun (_x : LocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F)) => (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) -> (Prod.{u2, u3} B F)) (LocalEquiv.hasCoeToFun.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F)) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i) p) (Prod.mk.{u2, u3} B F (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p) (FiberBundleCore.coordChange.{u1, u2, u3} ι B _inst_2 F _inst_3 Z (FiberBundleCore.indexAt.{u1, u2, u3} ι B _inst_2 F _inst_3 Z (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p)) i (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p) (Sigma.snd.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p))) but is expected to have type - forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_1 F _inst_2) (i : ι) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z), Eq.{max (succ u2) (succ u1)} (Prod.{u2, u1} B F) (LocalEquiv.toFun.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u1} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u3, u2, u1} ι B F _inst_1 _inst_2 Z i) p) (Prod.mk.{u2, u1} B F (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p) (FiberBundleCore.coordChange.{u3, u2, u1} ι B _inst_1 F _inst_2 Z (FiberBundleCore.indexAt.{u3, u2, u1} ι B _inst_1 F _inst_2 Z (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p)) i (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p) (Sigma.snd.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p))) + forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_2 F _inst_3) (i : ι) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z), Eq.{max (succ u2) (succ u1)} (Prod.{u2, u1} B F) (LocalEquiv.toFun.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u1} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u3, u2, u1} ι B F _inst_2 _inst_3 Z i) p) (Prod.mk.{u2, u1} B F (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p) (FiberBundleCore.coordChange.{u3, u2, u1} ι B _inst_2 F _inst_3 Z (FiberBundleCore.indexAt.{u3, u2, u1} ι B _inst_2 F _inst_3 Z (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p)) i (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p) (Sigma.snd.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_as_local_equiv_apply FiberBundleCore.localTrivAsLocalEquiv_applyₓ'. -/ theorem localTrivAsLocalEquiv_apply (p : Z.TotalSpace) : (Z.localTrivAsLocalEquiv i) p = ⟨p.1, Z.coordChange (Z.indexAt p.1) i p.1 p.2⟩ := @@ -606,9 +656,9 @@ theorem localTrivAsLocalEquiv_apply (p : Z.TotalSpace) : /- warning: fiber_bundle_core.local_triv_as_local_equiv_trans -> FiberBundleCore.localTrivAsLocalEquiv_trans is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι) (j : ι), HasEquivₓ.Equiv.{succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F)) (setoidHasEquiv.{succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F)) (LocalEquiv.eqOnSourceSetoid.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F))) (LocalEquiv.trans.{max u2 u3, max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (LocalEquiv.symm.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z j)) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (FiberBundleCore.trivChange.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i j)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι) (j : ι), HasEquivₓ.Equiv.{succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F)) (setoidHasEquiv.{succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F)) (LocalEquiv.eqOnSourceSetoid.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F))) (LocalEquiv.trans.{max u2 u3, max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (LocalEquiv.symm.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z j)) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Prod.{u2, u3} B F) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (FiberBundleCore.trivChange.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i j)) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι) (j : ι), HasEquiv.Equiv.{max (max (succ u3) (succ u2)) (succ (max u3 u2)), 0} (LocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F)) (instHasEquiv.{max (succ u3) (succ u2)} (LocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F)) (LocalEquiv.eqOnSourceSetoid.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F))) (LocalEquiv.trans.{max u3 u2, max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (LocalEquiv.symm.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z j)) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (FiberBundleCore.trivChange.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i j)) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι) (j : ι), HasEquiv.Equiv.{max (max (succ u3) (succ u2)) (succ (max u3 u2)), 0} (LocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F)) (instHasEquiv.{max (succ u3) (succ u2)} (LocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F)) (LocalEquiv.eqOnSourceSetoid.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F))) (LocalEquiv.trans.{max u3 u2, max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (LocalEquiv.symm.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z j)) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Prod.{u3, u2} B F) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (FiberBundleCore.trivChange.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i j)) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_as_local_equiv_trans FiberBundleCore.localTrivAsLocalEquiv_transₓ'. -/ /-- The composition of two local trivializations is the trivialization change Z.triv_change i j. -/ theorem localTrivAsLocalEquiv_trans (i j : ι) : @@ -640,9 +690,9 @@ variable (b : B) (a : F) /- warning: fiber_bundle_core.open_source' -> FiberBundleCore.open_source' is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι), IsOpen.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι), IsOpen.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι), IsOpen.{max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (LocalEquiv.source.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι), IsOpen.{max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (LocalEquiv.source.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.open_source' FiberBundleCore.open_source'ₓ'. -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem open_source' (i : ι) : IsOpen (Z.localTrivAsLocalEquiv i).source := @@ -714,9 +764,9 @@ def localTrivAt (b : B) : Trivialization F (π Z.Fiber) := /- warning: fiber_bundle_core.local_triv_at_def -> FiberBundleCore.localTrivAt_def is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (b : B), Eq.{max (succ u2) (succ u3) (succ (max u2 u3))} (Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z (FiberBundleCore.indexAt.{u1, u2, u3} ι B _inst_1 F _inst_2 Z b)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_1 _inst_2 Z b) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (b : B), Eq.{max (succ u2) (succ u3) (succ (max u2 u3))} (Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z (FiberBundleCore.indexAt.{u1, u2, u3} ι B _inst_2 F _inst_3 Z b)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_2 _inst_3 Z b) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (b : B), Eq.{max (succ u3) (succ u2)} (Trivialization.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z (FiberBundleCore.indexAt.{u1, u3, u2} ι B _inst_1 F _inst_2 Z b)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_1 _inst_2 Z b) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (b : B), Eq.{max (succ u3) (succ u2)} (Trivialization.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z (FiberBundleCore.indexAt.{u1, u3, u2} ι B _inst_2 F _inst_3 Z b)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_2 _inst_3 Z b) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_at_def FiberBundleCore.localTrivAt_defₓ'. -/ @[simp, mfld_simps] theorem localTrivAt_def (b : B) : Z.localTriv (Z.indexAt b) = Z.localTrivAt b := @@ -725,9 +775,9 @@ theorem localTrivAt_def (b : B) : Z.localTriv (Z.indexAt b) = Z.localTrivAt b := /- warning: fiber_bundle_core.continuous_const_section -> FiberBundleCore.continuous_const_section is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (v : F), (forall (i : ι) (j : ι) (x : B), (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) x (Inter.inter.{u2} (Set.{u2} B) (Set.hasInter.{u2} B) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_1 F _inst_2 Z i) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_1 F _inst_2 Z j))) -> (Eq.{succ u3} F (FiberBundleCore.coordChange.{u1, u2, u3} ι B _inst_1 F _inst_2 Z i j x v) v)) -> (Continuous.{u2, max u2 u3} B (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) ((fun (this : B -> (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) => this) (fun (x : B) => Sigma.mk.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) x v))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (v : F), (forall (i : ι) (j : ι) (x : B), (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) x (Inter.inter.{u2} (Set.{u2} B) (Set.hasInter.{u2} B) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_2 F _inst_3 Z i) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_2 F _inst_3 Z j))) -> (Eq.{succ u3} F (FiberBundleCore.coordChange.{u1, u2, u3} ι B _inst_2 F _inst_3 Z i j x v) v)) -> (Continuous.{u2, max u2 u3} B (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) ((fun (this : B -> (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) => this) (fun (x : B) => Sigma.mk.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) x v))) but is expected to have type - forall {ι : Type.{u2}} {B : Type.{u3}} {F : Type.{u1}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u2, u3, u1} ι B _inst_1 F _inst_2) (v : F), (forall (i : ι) (j : ι) (x : B), (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) x (Inter.inter.{u3} (Set.{u3} B) (Set.instInterSet.{u3} B) (FiberBundleCore.baseSet.{u2, u3, u1} ι B _inst_1 F _inst_2 Z i) (FiberBundleCore.baseSet.{u2, u3, u1} ι B _inst_1 F _inst_2 Z j))) -> (Eq.{succ u1} F (FiberBundleCore.coordChange.{u2, u3, u1} ι B _inst_1 F _inst_2 Z i j x v) v)) -> (Continuous.{u3, max u3 u1} B (FiberBundleCore.TotalSpace.{u2, u3, u1} ι B F _inst_1 _inst_2 Z) _inst_1 (FiberBundleCore.toTopologicalSpace.{u2, u3, u1} ι B F _inst_1 _inst_2 Z) ([mdata let_fun:1 (fun (this : B -> (FiberBundleCore.TotalSpace.{u2, u3, u1} ι B F _inst_1 _inst_2 Z)) => this) (fun (x : B) => Sigma.mk.{u3, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u2, u3, u1} ι B F _inst_1 _inst_2 Z x) x v)])) + forall {ι : Type.{u2}} {B : Type.{u3}} {F : Type.{u1}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u2, u3, u1} ι B _inst_2 F _inst_3) (v : F), (forall (i : ι) (j : ι) (x : B), (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) x (Inter.inter.{u3} (Set.{u3} B) (Set.instInterSet.{u3} B) (FiberBundleCore.baseSet.{u2, u3, u1} ι B _inst_2 F _inst_3 Z i) (FiberBundleCore.baseSet.{u2, u3, u1} ι B _inst_2 F _inst_3 Z j))) -> (Eq.{succ u1} F (FiberBundleCore.coordChange.{u2, u3, u1} ι B _inst_2 F _inst_3 Z i j x v) v)) -> (Continuous.{u3, max u3 u1} B (FiberBundleCore.TotalSpace.{u2, u3, u1} ι B F _inst_2 _inst_3 Z) _inst_2 (FiberBundleCore.toTopologicalSpace.{u2, u3, u1} ι B F _inst_2 _inst_3 Z) ([mdata let_fun:1 (fun (this : B -> (FiberBundleCore.TotalSpace.{u2, u3, u1} ι B F _inst_2 _inst_3 Z)) => this) (fun (x : B) => Sigma.mk.{u3, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u2, u3, u1} ι B F _inst_2 _inst_3 Z x) x v)])) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.continuous_const_section FiberBundleCore.continuous_const_sectionₓ'. -/ /-- If an element of `F` is invariant under all coordinate changes, then one can define a corresponding section of the fiber bundle, which is continuous. This applies in particular to the @@ -752,9 +802,9 @@ theorem continuous_const_section (v : F) /- warning: fiber_bundle_core.local_triv_as_local_equiv_coe -> FiberBundleCore.localTrivAsLocalEquiv_coe is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι), Eq.{succ (max u2 u3)} ((FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) -> (Prod.{u2, u3} B F)) (coeFn.{succ (max u2 u3), succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F)) (fun (_x : LocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F)) => (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) -> (Prod.{u2, u3} B F)) (LocalEquiv.hasCoeToFun.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F)) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)) (coeFn.{max (succ u2) (succ u3) (succ (max u2 u3)), max (succ (max u2 u3)) (succ u2) (succ u3)} (Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (fun (_x : Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) => (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) -> (Prod.{u2, u3} B F)) (Trivialization.hasCoeToFun.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι), Eq.{succ (max u2 u3)} ((FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) -> (Prod.{u2, u3} B F)) (coeFn.{succ (max u2 u3), succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F)) (fun (_x : LocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F)) => (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) -> (Prod.{u2, u3} B F)) (LocalEquiv.hasCoeToFun.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F)) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)) (coeFn.{max (succ u2) (succ u3) (succ (max u2 u3)), max (succ (max u2 u3)) (succ u2) (succ u3)} (Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (fun (_x : Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) => (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) -> (Prod.{u2, u3} B F)) (Trivialization.hasCoeToFun.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι), Eq.{max (succ u3) (succ u2)} ((FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) -> (Prod.{u3, u2} B F)) (LocalEquiv.toFun.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)) (Trivialization.toFun'.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι), Eq.{max (succ u3) (succ u2)} ((FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) -> (Prod.{u3, u2} B F)) (LocalEquiv.toFun.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)) (Trivialization.toFun'.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_as_local_equiv_coe FiberBundleCore.localTrivAsLocalEquiv_coeₓ'. -/ @[simp, mfld_simps] theorem localTrivAsLocalEquiv_coe : ⇑(Z.localTrivAsLocalEquiv i) = Z.localTriv i := @@ -763,9 +813,9 @@ theorem localTrivAsLocalEquiv_coe : ⇑(Z.localTrivAsLocalEquiv i) = Z.localTriv /- warning: fiber_bundle_core.local_triv_as_local_equiv_source -> FiberBundleCore.localTrivAsLocalEquiv_source is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι), Eq.{succ (max u2 u3)} (Set.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)) (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι), Eq.{succ (max u2 u3)} (Set.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)) (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)))) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι), Eq.{max (succ u3) (succ u2)} (Set.{max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (LocalEquiv.source.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)) (LocalEquiv.source.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)))) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι), Eq.{max (succ u3) (succ u2)} (Set.{max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (LocalEquiv.source.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)) (LocalEquiv.source.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_as_local_equiv_source FiberBundleCore.localTrivAsLocalEquiv_sourceₓ'. -/ @[simp, mfld_simps] theorem localTrivAsLocalEquiv_source : @@ -775,9 +825,9 @@ theorem localTrivAsLocalEquiv_source : /- warning: fiber_bundle_core.local_triv_as_local_equiv_target -> FiberBundleCore.localTrivAsLocalEquiv_target is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι), Eq.{succ (max u2 u3)} (Set.{max u2 u3} (Prod.{u2, u3} B F)) (LocalEquiv.target.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)) (LocalEquiv.target.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι), Eq.{succ (max u2 u3)} (Set.{max u2 u3} (Prod.{u2, u3} B F)) (LocalEquiv.target.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)) (LocalEquiv.target.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)))) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι), Eq.{max (succ u3) (succ u2)} (Set.{max u3 u2} (Prod.{u3, u2} B F)) (LocalEquiv.target.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)) (LocalEquiv.target.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)))) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι), Eq.{max (succ u3) (succ u2)} (Set.{max u3 u2} (Prod.{u3, u2} B F)) (LocalEquiv.target.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)) (LocalEquiv.target.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_as_local_equiv_target FiberBundleCore.localTrivAsLocalEquiv_targetₓ'. -/ @[simp, mfld_simps] theorem localTrivAsLocalEquiv_target : @@ -787,9 +837,9 @@ theorem localTrivAsLocalEquiv_target : /- warning: fiber_bundle_core.local_triv_as_local_equiv_symm -> FiberBundleCore.localTrivAsLocalEquiv_symm is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι), Eq.{succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (LocalEquiv.symm.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)) (LocalEquiv.symm.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι), Eq.{succ (max u2 u3)} (LocalEquiv.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (LocalEquiv.symm.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)) (LocalEquiv.symm.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)))) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι), Eq.{max (succ u3) (succ u2)} (LocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (LocalEquiv.symm.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)) (LocalEquiv.symm.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i)))) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι), Eq.{max (succ u3) (succ u2)} (LocalEquiv.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (LocalEquiv.symm.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.localTrivAsLocalEquiv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)) (LocalEquiv.symm.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i)))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_as_local_equiv_symm FiberBundleCore.localTrivAsLocalEquiv_symmₓ'. -/ @[simp, mfld_simps] theorem localTrivAsLocalEquiv_symm : @@ -799,9 +849,9 @@ theorem localTrivAsLocalEquiv_symm : /- warning: fiber_bundle_core.base_set_at -> FiberBundleCore.baseSet_at is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι), Eq.{succ u2} (Set.{u2} B) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_1 F _inst_2 Z i) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι), Eq.{succ u2} (Set.{u2} B) (FiberBundleCore.baseSet.{u1, u2, u3} ι B _inst_2 F _inst_3 Z i) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i)) but is expected to have type - forall {ι : Type.{u2}} {B : Type.{u3}} {F : Type.{u1}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u2, u3, u1} ι B _inst_1 F _inst_2) (i : ι), Eq.{succ u3} (Set.{u3} B) (FiberBundleCore.baseSet.{u2, u3, u1} ι B _inst_1 F _inst_2 Z i) (Trivialization.baseSet.{u3, u1, max u3 u1} B F (FiberBundleCore.TotalSpace.{u2, u3, u1} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u2, u3, u1} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u2, u3, u1} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u2, u3, u1} ι B F _inst_1 _inst_2 Z i)) + forall {ι : Type.{u2}} {B : Type.{u3}} {F : Type.{u1}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u2, u3, u1} ι B _inst_2 F _inst_3) (i : ι), Eq.{succ u3} (Set.{u3} B) (FiberBundleCore.baseSet.{u2, u3, u1} ι B _inst_2 F _inst_3 Z i) (Trivialization.baseSet.{u3, u1, max u3 u1} B F (FiberBundleCore.TotalSpace.{u2, u3, u1} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u2, u3, u1} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u2, u3, u1} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u2, u3, u1} ι B F _inst_2 _inst_3 Z i)) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.base_set_at FiberBundleCore.baseSet_atₓ'. -/ @[simp, mfld_simps] theorem baseSet_at : Z.baseSet i = (Z.localTriv i).baseSet := @@ -810,9 +860,9 @@ theorem baseSet_at : Z.baseSet i = (Z.localTriv i).baseSet := /- warning: fiber_bundle_core.local_triv_apply -> FiberBundleCore.localTriv_apply is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z), Eq.{max (succ u2) (succ u3)} (Prod.{u2, u3} B F) (coeFn.{max (succ u2) (succ u3) (succ (max u2 u3)), max (succ (max u2 u3)) (succ u2) (succ u3)} (Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (fun (_x : Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) => (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) -> (Prod.{u2, u3} B F)) (Trivialization.hasCoeToFun.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i) p) (Prod.mk.{u2, u3} B F (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p) (FiberBundleCore.coordChange.{u1, u2, u3} ι B _inst_1 F _inst_2 Z (FiberBundleCore.indexAt.{u1, u2, u3} ι B _inst_1 F _inst_2 Z (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p)) i (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p) (Sigma.snd.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z), Eq.{max (succ u2) (succ u3)} (Prod.{u2, u3} B F) (coeFn.{max (succ u2) (succ u3) (succ (max u2 u3)), max (succ (max u2 u3)) (succ u2) (succ u3)} (Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (fun (_x : Trivialization.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) => (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) -> (Prod.{u2, u3} B F)) (Trivialization.hasCoeToFun.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i) p) (Prod.mk.{u2, u3} B F (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p) (FiberBundleCore.coordChange.{u1, u2, u3} ι B _inst_2 F _inst_3 Z (FiberBundleCore.indexAt.{u1, u2, u3} ι B _inst_2 F _inst_3 Z (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p)) i (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p) (Sigma.snd.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p))) but is expected to have type - forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_1 F _inst_2) (i : ι) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z), Eq.{max (succ u2) (succ u1)} (Prod.{u2, u1} B F) (Trivialization.toFun'.{u2, u1, max u2 u1} B F (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.proj.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u3, u2, u1} ι B F _inst_1 _inst_2 Z i) p) (Prod.mk.{u2, u1} B F (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p) (FiberBundleCore.coordChange.{u3, u2, u1} ι B _inst_1 F _inst_2 Z (FiberBundleCore.indexAt.{u3, u2, u1} ι B _inst_1 F _inst_2 Z (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p)) i (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p) (Sigma.snd.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p))) + forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_2 F _inst_3) (i : ι) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z), Eq.{max (succ u2) (succ u1)} (Prod.{u2, u1} B F) (Trivialization.toFun'.{u2, u1, max u2 u1} B F (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.proj.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u3, u2, u1} ι B F _inst_2 _inst_3 Z i) p) (Prod.mk.{u2, u1} B F (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p) (FiberBundleCore.coordChange.{u3, u2, u1} ι B _inst_2 F _inst_3 Z (FiberBundleCore.indexAt.{u3, u2, u1} ι B _inst_2 F _inst_3 Z (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p)) i (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p) (Sigma.snd.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_apply FiberBundleCore.localTriv_applyₓ'. -/ @[simp, mfld_simps] theorem localTriv_apply (p : Z.TotalSpace) : @@ -822,9 +872,9 @@ theorem localTriv_apply (p : Z.TotalSpace) : /- warning: fiber_bundle_core.local_triv_at_apply -> FiberBundleCore.localTrivAt_apply is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z), Eq.{max (succ u2) (succ u3)} (Prod.{u2, u3} B F) (coeFn.{max (succ u2) (succ u3) (succ (max u2 u3)), max (succ (max u2 u3)) (succ u2) (succ u3)} (Trivialization.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z))) (fun (_x : Trivialization.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z))) => (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) -> (Prod.{u2, u3} B F)) (Trivialization.hasCoeToFun.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_1 _inst_2 Z (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p)) p) (Prod.mk.{u2, u3} B F (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p) (Sigma.snd.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z), Eq.{max (succ u2) (succ u3)} (Prod.{u2, u3} B F) (coeFn.{max (succ u2) (succ u3) (succ (max u2 u3)), max (succ (max u2 u3)) (succ u2) (succ u3)} (Trivialization.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z))) (fun (_x : Trivialization.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z))) => (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) -> (Prod.{u2, u3} B F)) (Trivialization.hasCoeToFun.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_2 _inst_3 Z (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p)) p) (Prod.mk.{u2, u3} B F (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p) (Sigma.snd.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p)) but is expected to have type - forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_1 F _inst_2) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z), Eq.{max (succ u2) (succ u1)} (Prod.{u2, u1} B F) (Trivialization.toFun'.{u2, u1, max u2 u1} B F (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTrivAt.{u3, u2, u1} ι B F _inst_1 _inst_2 Z (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p)) p) (Prod.mk.{u2, u1} B F (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p) (Sigma.snd.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p)) + forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_2 F _inst_3) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z), Eq.{max (succ u2) (succ u1)} (Prod.{u2, u1} B F) (Trivialization.toFun'.{u2, u1, max u2 u1} B F (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTrivAt.{u3, u2, u1} ι B F _inst_2 _inst_3 Z (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p)) p) (Prod.mk.{u2, u1} B F (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p) (Sigma.snd.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p)) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_at_apply FiberBundleCore.localTrivAt_applyₓ'. -/ @[simp, mfld_simps] theorem localTrivAt_apply (p : Z.TotalSpace) : (Z.localTrivAt p.1) p = ⟨p.1, p.2⟩ := @@ -835,9 +885,9 @@ theorem localTrivAt_apply (p : Z.TotalSpace) : (Z.localTrivAt p.1) p = ⟨p.1, p /- warning: fiber_bundle_core.local_triv_at_apply_mk -> FiberBundleCore.localTrivAt_apply_mk is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (b : B) (a : F), Eq.{max (succ u2) (succ u3)} (Prod.{u2, u3} B F) (coeFn.{max (succ u2) (succ u3) (succ (max u2 u3)), max (succ (max u2 u3)) (succ u2) (succ u3)} (Trivialization.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z))) (fun (_x : Trivialization.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z))) => (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) -> (Prod.{u2, u3} B F)) (Trivialization.hasCoeToFun.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_1 _inst_2 Z b) (Sigma.mk.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) b a)) (Prod.mk.{u2, u3} B F b a) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (b : B) (a : F), Eq.{max (succ u2) (succ u3)} (Prod.{u2, u3} B F) (coeFn.{max (succ u2) (succ u3) (succ (max u2 u3)), max (succ (max u2 u3)) (succ u2) (succ u3)} (Trivialization.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z))) (fun (_x : Trivialization.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z))) => (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) -> (Prod.{u2, u3} B F)) (Trivialization.hasCoeToFun.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_2 _inst_3 Z b) (Sigma.mk.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) b a)) (Prod.mk.{u2, u3} B F b a) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (b : B) (a : F), Eq.{max (succ u3) (succ u2)} (Prod.{u3, u2} B F) (Trivialization.toFun'.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_1 _inst_2 Z b) (Sigma.mk.{u3, u2} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z x) b a)) (Prod.mk.{u3, u2} B F b a) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (b : B) (a : F), Eq.{max (succ u3) (succ u2)} (Prod.{u3, u2} B F) (Trivialization.toFun'.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_2 _inst_3 Z b) (Sigma.mk.{u3, u2} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z x) b a)) (Prod.mk.{u3, u2} B F b a) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_at_apply_mk FiberBundleCore.localTrivAt_apply_mkₓ'. -/ @[simp, mfld_simps] theorem localTrivAt_apply_mk (b : B) (a : F) : (Z.localTrivAt b) ⟨b, a⟩ = ⟨b, a⟩ := @@ -846,9 +896,9 @@ theorem localTrivAt_apply_mk (b : B) (a : F) : (Z.localTrivAt b) ⟨b, a⟩ = /- warning: fiber_bundle_core.mem_local_triv_source -> FiberBundleCore.mem_localTriv_source is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z), Iff (Membership.Mem.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Set.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (Set.hasMem.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) p (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i))))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z), Iff (Membership.Mem.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Set.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (Set.hasMem.{max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) p (LocalEquiv.source.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i))))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i))) but is expected to have type - forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_1 F _inst_2) (i : ι) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z), Iff (Membership.mem.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (Set.{max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) (Set.instMembershipSet.{max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) p (LocalEquiv.source.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u1} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u1} B F) (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u2, u1} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u1, max u2 u1} B F (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u3, u2, u1} ι B F _inst_1 _inst_2 Z i))))) (Membership.mem.{u2, u2} B (Set.{u2} B) (Set.instMembershipSet.{u2} B) (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p) (Trivialization.baseSet.{u2, u1, max u2 u1} B F (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u3, u2, u1} ι B F _inst_1 _inst_2 Z i))) + forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_2 F _inst_3) (i : ι) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z), Iff (Membership.mem.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (Set.{max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) (Set.instMembershipSet.{max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) p (LocalEquiv.source.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u1} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u1} B F) (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u2, u1} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u1, max u2 u1} B F (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u3, u2, u1} ι B F _inst_2 _inst_3 Z i))))) (Membership.mem.{u2, u2} B (Set.{u2} B) (Set.instMembershipSet.{u2} B) (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p) (Trivialization.baseSet.{u2, u1, max u2 u1} B F (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u3, u2, u1} ι B F _inst_2 _inst_3 Z i))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.mem_local_triv_source FiberBundleCore.mem_localTriv_sourceₓ'. -/ @[simp, mfld_simps] theorem mem_localTriv_source (p : Z.TotalSpace) : @@ -858,9 +908,9 @@ theorem mem_localTriv_source (p : Z.TotalSpace) : /- warning: fiber_bundle_core.mem_local_triv_at_source -> FiberBundleCore.mem_localTrivAt_source is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (b : B), Iff (Membership.Mem.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Set.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z))) (Set.hasMem.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z))) p (LocalEquiv.source.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_1 _inst_2 Z b))))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) p) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_1 _inst_2 Z b))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (p : FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (b : B), Iff (Membership.Mem.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Set.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z))) (Set.hasMem.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z))) p (LocalEquiv.source.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_2 _inst_3 Z b))))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Sigma.fst.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) p) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_2 _inst_3 Z b))) but is expected to have type - forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_1 F _inst_2) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (b : B), Iff (Membership.mem.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (Set.{max u2 u1} (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z))) (Set.instMembershipSet.{max u2 u1} (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z))) p (LocalEquiv.source.{max u2 u1, max u2 u1} (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) (Prod.{u2, u1} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u1, max u2 u1} (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) (Prod.{u2, u1} B F) (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u2, u1} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u1, max u2 u1} B F (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u3, u2, u1} ι B F _inst_1 _inst_2 Z b))))) (Membership.mem.{u2, u2} B (Set.{u2} B) (Set.instMembershipSet.{u2} B) (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z x) p) (Trivialization.baseSet.{u2, u1, max u2 u1} B F (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u3, u2, u1} ι B F _inst_1 _inst_2 Z b))) + forall {ι : Type.{u3}} {B : Type.{u2}} {F : Type.{u1}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (Z : FiberBundleCore.{u3, u2, u1} ι B _inst_2 F _inst_3) (p : FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (b : B), Iff (Membership.mem.{max u2 u1, max u2 u1} (FiberBundleCore.TotalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (Set.{max u2 u1} (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z))) (Set.instMembershipSet.{max u2 u1} (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z))) p (LocalEquiv.source.{max u2 u1, max u2 u1} (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) (Prod.{u2, u1} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u1, max u2 u1} (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) (Prod.{u2, u1} B F) (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u2, u1} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u1, max u2 u1} B F (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u3, u2, u1} ι B F _inst_2 _inst_3 Z b))))) (Membership.mem.{u2, u2} B (Set.{u2} B) (Set.instMembershipSet.{u2} B) (Sigma.fst.{u2, u1} B (fun (x : B) => FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z x) p) (Trivialization.baseSet.{u2, u1, max u2 u1} B F (Bundle.TotalSpace.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u3, u2, u1} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u1} B (FiberBundleCore.Fiber.{u3, u2, u1} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u3, u2, u1} ι B F _inst_2 _inst_3 Z b))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.mem_local_triv_at_source FiberBundleCore.mem_localTrivAt_sourceₓ'. -/ @[simp, mfld_simps] theorem mem_localTrivAt_source (p : Z.TotalSpace) (b : B) : @@ -870,9 +920,9 @@ theorem mem_localTrivAt_source (p : Z.TotalSpace) (b : B) : /- warning: fiber_bundle_core.mem_source_at -> FiberBundleCore.mem_source_at is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (b : B) (a : F), Membership.Mem.{max u2 u3, max u2 u3} (Sigma.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x)) (Set.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z))) (Set.hasMem.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z))) (Sigma.mk.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) b a) (LocalEquiv.source.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_1 _inst_2 Z b)))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (b : B) (a : F), Membership.Mem.{max u2 u3, max u2 u3} (Sigma.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x)) (Set.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z))) (Set.hasMem.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z))) (Sigma.mk.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) b a) (LocalEquiv.source.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_2 _inst_3 Z b)))) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (b : B) (a : F), Membership.mem.{max u3 u2, max u3 u2} (Sigma.{u3, u2} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z x)) (Set.{max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z))) (Set.instMembershipSet.{max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z))) (Sigma.mk.{u3, u2} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z x) b a) (LocalEquiv.source.{max u3 u2, max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_1 _inst_2 Z b)))) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (b : B) (a : F), Membership.mem.{max u3 u2, max u3 u2} (Sigma.{u3, u2} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z x)) (Set.{max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z))) (Set.instMembershipSet.{max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z))) (Sigma.mk.{u3, u2} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z x) b a) (LocalEquiv.source.{max u3 u2, max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_2 _inst_3 Z b)))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.mem_source_at FiberBundleCore.mem_source_atₓ'. -/ @[simp, mfld_simps] theorem mem_source_at : (⟨b, a⟩ : Z.TotalSpace) ∈ (Z.localTrivAt b).source := @@ -883,9 +933,9 @@ theorem mem_source_at : (⟨b, a⟩ : Z.TotalSpace) ∈ (Z.localTrivAt b).source /- warning: fiber_bundle_core.mem_local_triv_target -> FiberBundleCore.mem_localTriv_target is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι) (p : Prod.{u2, u3} B F), Iff (Membership.Mem.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Set.{max u2 u3} (Prod.{u2, u3} B F)) (Set.hasMem.{max u2 u3} (Prod.{u2, u3} B F)) p (LocalEquiv.target.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i))))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Prod.fst.{u2, u3} B F p) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι) (p : Prod.{u2, u3} B F), Iff (Membership.Mem.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Set.{max u2 u3} (Prod.{u2, u3} B F)) (Set.hasMem.{max u2 u3} (Prod.{u2, u3} B F)) p (LocalEquiv.target.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i))))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Prod.fst.{u2, u3} B F p) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i))) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι) (p : Prod.{u3, u2} B F), Iff (Membership.mem.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instMembershipSet.{max u3 u2} (Prod.{u3, u2} B F)) p (LocalEquiv.target.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i))))) (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) (Prod.fst.{u3, u2} B F p) (Trivialization.baseSet.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i))) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι) (p : Prod.{u3, u2} B F), Iff (Membership.mem.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instMembershipSet.{max u3 u2} (Prod.{u3, u2} B F)) p (LocalEquiv.target.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i))))) (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) (Prod.fst.{u3, u2} B F p) (Trivialization.baseSet.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.mem_local_triv_target FiberBundleCore.mem_localTriv_targetₓ'. -/ @[simp, mfld_simps] theorem mem_localTriv_target (p : B × F) : @@ -895,9 +945,9 @@ theorem mem_localTriv_target (p : B × F) : /- warning: fiber_bundle_core.mem_local_triv_at_target -> FiberBundleCore.mem_localTrivAt_target is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (p : Prod.{u2, u3} B F) (b : B), Iff (Membership.Mem.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Set.{max u2 u3} (Prod.{u2, u3} B F)) (Set.hasMem.{max u2 u3} (Prod.{u2, u3} B F)) p (LocalEquiv.target.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_1 _inst_2 Z b))))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Prod.fst.{u2, u3} B F p) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_1 _inst_2 Z b))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (p : Prod.{u2, u3} B F) (b : B), Iff (Membership.Mem.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (Set.{max u2 u3} (Prod.{u2, u3} B F)) (Set.hasMem.{max u2 u3} (Prod.{u2, u3} B F)) p (LocalEquiv.target.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (Prod.{u2, u3} B F) (LocalHomeomorph.toLocalEquiv.{max u2 u3, max u2 u3} (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_2 _inst_3 Z b))))) (Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) (Prod.fst.{u2, u3} B F p) (Trivialization.baseSet.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_2 _inst_3 Z b))) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (p : Prod.{u3, u2} B F) (b : B), Iff (Membership.mem.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instMembershipSet.{max u3 u2} (Prod.{u3, u2} B F)) p (LocalEquiv.target.{max u3 u2, max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_1 _inst_2 Z b))))) (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) (Prod.fst.{u3, u2} B F p) (Trivialization.baseSet.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_1 _inst_2 Z b))) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (p : Prod.{u3, u2} B F) (b : B), Iff (Membership.mem.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instMembershipSet.{max u3 u2} (Prod.{u3, u2} B F)) p (LocalEquiv.target.{max u3 u2, max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (Prod.{u3, u2} B F) (LocalHomeomorph.toLocalEquiv.{max u3 u2, max u3 u2} (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_2 _inst_3 Z b))))) (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) (Prod.fst.{u3, u2} B F p) (Trivialization.baseSet.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_2 _inst_3 Z b))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.mem_local_triv_at_target FiberBundleCore.mem_localTrivAt_targetₓ'. -/ @[simp, mfld_simps] theorem mem_localTrivAt_target (p : B × F) (b : B) : @@ -907,9 +957,9 @@ theorem mem_localTrivAt_target (p : B × F) (b : B) : /- warning: fiber_bundle_core.local_triv_symm_apply -> FiberBundleCore.localTriv_symm_apply is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (i : ι) (p : Prod.{u2, u3} B F), Eq.{max (succ u2) (succ u3)} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (coeFn.{succ (max u2 u3), succ (max u2 u3)} (LocalHomeomorph.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (fun (_x : LocalHomeomorph.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) => (Prod.{u2, u3} B F) -> (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (LocalHomeomorph.hasCoeToFun.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (LocalHomeomorph.symm.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_1 _inst_2 Z i))) p) (Sigma.mk.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z x) (Prod.fst.{u2, u3} B F p) (FiberBundleCore.coordChange.{u1, u2, u3} ι B _inst_1 F _inst_2 Z i (FiberBundleCore.indexAt.{u1, u2, u3} ι B _inst_1 F _inst_2 Z (Prod.fst.{u2, u3} B F p)) (Prod.fst.{u2, u3} B F p) (Prod.snd.{u2, u3} B F p))) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (i : ι) (p : Prod.{u2, u3} B F), Eq.{max (succ u2) (succ u3)} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (coeFn.{succ (max u2 u3), succ (max u2 u3)} (LocalHomeomorph.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (fun (_x : LocalHomeomorph.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) => (Prod.{u2, u3} B F) -> (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (LocalHomeomorph.hasCoeToFun.{max u2 u3, max u2 u3} (Prod.{u2, u3} B F) (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (LocalHomeomorph.symm.{max u2 u3, max u2 u3} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.{u2, u3} B F) (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Prod.topologicalSpace.{u2, u3} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u2, u3, max u2 u3} B F (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u2, u3} ι B F _inst_2 _inst_3 Z i))) p) (Sigma.mk.{u2, u3} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z x) (Prod.fst.{u2, u3} B F p) (FiberBundleCore.coordChange.{u1, u2, u3} ι B _inst_2 F _inst_3 Z i (FiberBundleCore.indexAt.{u1, u2, u3} ι B _inst_2 F _inst_3 Z (Prod.fst.{u2, u3} B F p)) (Prod.fst.{u2, u3} B F p) (Prod.snd.{u2, u3} B F p))) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (i : ι) (p : Prod.{u3, u2} B F), Eq.{max (succ u3) (succ u2)} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (LocalHomeomorph.toFun'.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (LocalHomeomorph.symm.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_1 _inst_2 Z i))) p) (Sigma.mk.{u3, u2} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z x) (Prod.fst.{u3, u2} B F p) (FiberBundleCore.coordChange.{u1, u3, u2} ι B _inst_1 F _inst_2 Z i (FiberBundleCore.indexAt.{u1, u3, u2} ι B _inst_1 F _inst_2 Z (Prod.fst.{u3, u2} B F p)) (Prod.fst.{u3, u2} B F p) (Prod.snd.{u3, u2} B F p))) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (i : ι) (p : Prod.{u3, u2} B F), Eq.{max (succ u3) (succ u2)} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (LocalHomeomorph.toFun'.{max u3 u2, max u3 u2} (Prod.{u3, u2} B F) (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (LocalHomeomorph.symm.{max u3 u2, max u3 u2} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Prod.{u3, u2} B F) (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (Trivialization.toLocalHomeomorph.{u3, u2, max u3 u2} B F (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (FiberBundleCore.localTriv.{u1, u3, u2} ι B F _inst_2 _inst_3 Z i))) p) (Sigma.mk.{u3, u2} B (fun (x : B) => FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z x) (Prod.fst.{u3, u2} B F p) (FiberBundleCore.coordChange.{u1, u3, u2} ι B _inst_2 F _inst_3 Z i (FiberBundleCore.indexAt.{u1, u3, u2} ι B _inst_2 F _inst_3 Z (Prod.fst.{u3, u2} B F p)) (Prod.fst.{u3, u2} B F p) (Prod.snd.{u3, u2} B F p))) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.local_triv_symm_apply FiberBundleCore.localTriv_symm_applyₓ'. -/ @[simp, mfld_simps] theorem localTriv_symm_apply (p : B × F) : @@ -919,9 +969,9 @@ theorem localTriv_symm_apply (p : B × F) : /- warning: fiber_bundle_core.mem_local_triv_at_base_set -> FiberBundleCore.mem_localTrivAt_baseSet is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2) (b : B), Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) b (Trivialization.baseSet.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_1 _inst_2 Z b)) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3) (b : B), Membership.Mem.{u2, u2} B (Set.{u2} B) (Set.hasMem.{u2} B) b (Trivialization.baseSet.{u2, u3, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u2, u3} B (FiberBundleCore.Fiber.{u1, u2, u3} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u2, u3} ι B F _inst_2 _inst_3 Z b)) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2) (b : B), Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) b (Trivialization.baseSet.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) _inst_1 _inst_2 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_1 _inst_2 Z)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_1 _inst_2 Z b)) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3) (b : B), Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) b (Trivialization.baseSet.{u3, u2, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) _inst_2 _inst_3 (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) (Bundle.TotalSpace.proj.{u3, u2} B (FiberBundleCore.Fiber.{u1, u3, u2} ι B F _inst_2 _inst_3 Z)) (FiberBundleCore.localTrivAt.{u1, u3, u2} ι B F _inst_2 _inst_3 Z b)) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.mem_local_triv_at_base_set FiberBundleCore.mem_localTrivAt_baseSetₓ'. -/ @[simp, mfld_simps] theorem mem_localTrivAt_baseSet (b : B) : b ∈ (Z.localTrivAt b).baseSet := @@ -995,9 +1045,9 @@ instance fiberBundle : FiberBundle F Z.Fiber /- warning: fiber_bundle_core.continuous_proj -> FiberBundleCore.continuous_proj is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2), Continuous.{max u2 u3, u2} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) B (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3), Continuous.{max u2 u3, u2} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) B (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2), Continuous.{max u3 u2, u3} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) B (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3), Continuous.{max u3 u2, u3} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) B (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.continuous_proj FiberBundleCore.continuous_projₓ'. -/ /-- The projection on the base of a fiber bundle created from core is continuous -/ theorem continuous_proj : Continuous Z.proj := @@ -1006,9 +1056,9 @@ theorem continuous_proj : Continuous Z.proj := /- warning: fiber_bundle_core.is_open_map_proj -> FiberBundleCore.isOpenMap_proj is a dubious translation: lean 3 declaration is - forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_1 F _inst_2), IsOpenMap.{max u2 u3, u2} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) B (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) _inst_1 (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_1 _inst_2 Z) + forall {ι : Type.{u1}} {B : Type.{u2}} {F : Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u3} F] (Z : FiberBundleCore.{u1, u2, u3} ι B _inst_2 F _inst_3), IsOpenMap.{max u2 u3, u2} (FiberBundleCore.TotalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) B (FiberBundleCore.toTopologicalSpace.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) _inst_2 (FiberBundleCore.proj.{u1, u2, u3} ι B F _inst_2 _inst_3 Z) but is expected to have type - forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_1 F _inst_2), IsOpenMap.{max u3 u2, u3} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) B (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) _inst_1 (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_1 _inst_2 Z) + forall {ι : Type.{u1}} {B : Type.{u3}} {F : Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (Z : FiberBundleCore.{u1, u3, u2} ι B _inst_2 F _inst_3), IsOpenMap.{max u3 u2, u3} (FiberBundleCore.TotalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) B (FiberBundleCore.toTopologicalSpace.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) _inst_2 (FiberBundleCore.proj.{u1, u3, u2} ι B F _inst_2 _inst_3 Z) Case conversion may be inaccurate. Consider using '#align fiber_bundle_core.is_open_map_proj FiberBundleCore.isOpenMap_projₓ'. -/ /-- The projection on the base of a fiber bundle created from core is an open map -/ theorem isOpenMap_proj : IsOpenMap Z.proj := @@ -1054,9 +1104,9 @@ def totalSpaceTopology (a : FiberPrebundle F E) : TopologicalSpace (TotalSpace E /- warning: fiber_prebundle.continuous_symm_of_mem_pretrivialization_atlas -> FiberPrebundle.continuous_symm_of_mem_pretrivializationAtlas is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_1 _inst_2) {e : Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E)}, (Membership.Mem.{max u1 u2 u1 u3, max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E)) (Set.{max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E))) (Set.hasMem.{max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E))) e (FiberPrebundle.pretrivializationAtlas.{u1, u2, u3} B F E _inst_1 _inst_2 a)) -> (ContinuousOn.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E) (Prod.topologicalSpace.{u1, u2} B F _inst_1 _inst_2) (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_1 _inst_2 a) (coeFn.{max (succ (max u1 u2)) (succ (max u1 u3)), max (succ (max u1 u2)) (succ (max u1 u3))} (LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (fun (_x : LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) => (Prod.{u1, u2} B F) -> (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.hasCoeToFun.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.symm.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) e))) (LocalEquiv.target.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) e))) + forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_2 _inst_3) {e : Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E)}, (Membership.Mem.{max u1 u2 u1 u3, max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E)) (Set.{max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E))) (Set.hasMem.{max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E))) e (FiberPrebundle.pretrivializationAtlas.{u1, u2, u3} B F E _inst_2 _inst_3 a)) -> (ContinuousOn.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E) (Prod.topologicalSpace.{u1, u2} B F _inst_2 _inst_3) (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_2 _inst_3 a) (coeFn.{max (succ (max u1 u2)) (succ (max u1 u3)), max (succ (max u1 u2)) (succ (max u1 u3))} (LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (fun (_x : LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) => (Prod.{u1, u2} B F) -> (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.hasCoeToFun.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.symm.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) e))) (LocalEquiv.target.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) e))) but is expected to have type - forall {B : Type.{u3}} {F : Type.{u2}} {E : B -> Type.{u1}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u3, u2, u1} B F E _inst_1 _inst_2) {e : Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E)}, (Membership.mem.{max (max u3 u2) u1, max (max u3 u2) u1} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E)) (Set.{max (max (max u3 u1) u2) u3} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E))) (Set.instMembershipSet.{max (max u3 u2) u1} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E))) e (FiberPrebundle.pretrivializationAtlas.{u3, u2, u1} B F E _inst_1 _inst_2 a)) -> (ContinuousOn.{max u3 u2, max u3 u1} (Prod.{u3, u2} B F) (Bundle.TotalSpace.{u3, u1} B E) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (FiberPrebundle.totalSpaceTopology.{u3, u2, u1} B F E _inst_1 _inst_2 a) (LocalEquiv.toFun.{max u3 u2, max u3 u1} (Prod.{u3, u2} B F) (Bundle.TotalSpace.{u3, u1} B E) (LocalEquiv.symm.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E) e))) (LocalEquiv.target.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E) e))) + forall {B : Type.{u3}} {F : Type.{u2}} {E : B -> Type.{u1}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u3, u2, u1} B F E _inst_2 _inst_3) {e : Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E)}, (Membership.mem.{max (max u3 u2) u1, max (max u3 u2) u1} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E)) (Set.{max (max (max u3 u1) u2) u3} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E))) (Set.instMembershipSet.{max (max u3 u2) u1} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E))) e (FiberPrebundle.pretrivializationAtlas.{u3, u2, u1} B F E _inst_2 _inst_3 a)) -> (ContinuousOn.{max u3 u2, max u3 u1} (Prod.{u3, u2} B F) (Bundle.TotalSpace.{u3, u1} B E) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (FiberPrebundle.totalSpaceTopology.{u3, u2, u1} B F E _inst_2 _inst_3 a) (LocalEquiv.toFun.{max u3 u2, max u3 u1} (Prod.{u3, u2} B F) (Bundle.TotalSpace.{u3, u1} B E) (LocalEquiv.symm.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E) e))) (LocalEquiv.target.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E) e))) Case conversion may be inaccurate. Consider using '#align fiber_prebundle.continuous_symm_of_mem_pretrivialization_atlas FiberPrebundle.continuous_symm_of_mem_pretrivializationAtlasₓ'. -/ theorem continuous_symm_of_mem_pretrivializationAtlas (he : e ∈ a.pretrivializationAtlas) : @ContinuousOn _ _ _ a.totalSpaceTopology e.toLocalEquiv.symm e.target := @@ -1069,9 +1119,9 @@ theorem continuous_symm_of_mem_pretrivializationAtlas (he : e ∈ a.pretrivializ /- warning: fiber_prebundle.is_open_source -> FiberPrebundle.isOpen_source is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_1 _inst_2) (e : Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E)), IsOpen.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E) (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_1 _inst_2 a) (LocalEquiv.source.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) e)) + forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_2 _inst_3) (e : Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E)), IsOpen.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E) (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_2 _inst_3 a) (LocalEquiv.source.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) e)) but is expected to have type - forall {B : Type.{u3}} {F : Type.{u2}} {E : B -> Type.{u1}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u3, u2, u1} B F E _inst_1 _inst_2) (e : Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E)), IsOpen.{max u3 u1} (Bundle.TotalSpace.{u3, u1} B E) (FiberPrebundle.totalSpaceTopology.{u3, u2, u1} B F E _inst_1 _inst_2 a) (LocalEquiv.source.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E) e)) + forall {B : Type.{u3}} {F : Type.{u2}} {E : B -> Type.{u1}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u3, u2, u1} B F E _inst_2 _inst_3) (e : Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E)), IsOpen.{max u3 u1} (Bundle.TotalSpace.{u3, u1} B E) (FiberPrebundle.totalSpaceTopology.{u3, u2, u1} B F E _inst_2 _inst_3 a) (LocalEquiv.source.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E) e)) Case conversion may be inaccurate. Consider using '#align fiber_prebundle.is_open_source FiberPrebundle.isOpen_sourceₓ'. -/ theorem isOpen_source (e : Pretrivialization F (π E)) : is_open[a.totalSpaceTopology] e.source := by @@ -1086,9 +1136,9 @@ theorem isOpen_source (e : Pretrivialization F (π E)) : is_open[a.totalSpaceTop /- warning: fiber_prebundle.is_open_target_of_mem_pretrivialization_atlas_inter -> FiberPrebundle.isOpen_target_of_mem_pretrivializationAtlas_inter is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_1 _inst_2) (e : Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E)) (e' : Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E)), (Membership.Mem.{max u1 u2 u1 u3, max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E)) (Set.{max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E))) (Set.hasMem.{max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E))) e' (FiberPrebundle.pretrivializationAtlas.{u1, u2, u3} B F E _inst_1 _inst_2 a)) -> (IsOpen.{max u1 u2} (Prod.{u1, u2} B F) (Prod.topologicalSpace.{u1, u2} B F _inst_1 _inst_2) (Inter.inter.{max u1 u2} (Set.{max u1 u2} (Prod.{u1, u2} B F)) (Set.hasInter.{max u1 u2} (Prod.{u1, u2} B F)) (LocalEquiv.target.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) e')) (Set.preimage.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E) (coeFn.{max (succ (max u1 u2)) (succ (max u1 u3)), max (succ (max u1 u2)) (succ (max u1 u3))} (LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (fun (_x : LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) => (Prod.{u1, u2} B F) -> (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.hasCoeToFun.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.symm.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) e'))) (LocalEquiv.source.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) e))))) + forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_2 _inst_3) (e : Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E)) (e' : Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E)), (Membership.Mem.{max u1 u2 u1 u3, max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E)) (Set.{max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E))) (Set.hasMem.{max u1 u2 u1 u3} (Pretrivialization.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E))) e' (FiberPrebundle.pretrivializationAtlas.{u1, u2, u3} B F E _inst_2 _inst_3 a)) -> (IsOpen.{max u1 u2} (Prod.{u1, u2} B F) (Prod.topologicalSpace.{u1, u2} B F _inst_2 _inst_3) (Inter.inter.{max u1 u2} (Set.{max u1 u2} (Prod.{u1, u2} B F)) (Set.hasInter.{max u1 u2} (Prod.{u1, u2} B F)) (LocalEquiv.target.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) e')) (Set.preimage.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E) (coeFn.{max (succ (max u1 u2)) (succ (max u1 u3)), max (succ (max u1 u2)) (succ (max u1 u3))} (LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (fun (_x : LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) => (Prod.{u1, u2} B F) -> (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.hasCoeToFun.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.symm.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) e'))) (LocalEquiv.source.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) e))))) but is expected to have type - forall {B : Type.{u3}} {F : Type.{u2}} {E : B -> Type.{u1}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u3, u2, u1} B F E _inst_1 _inst_2) (e : Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E)) (e' : Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E)), (Membership.mem.{max (max u3 u2) u1, max (max u3 u2) u1} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E)) (Set.{max (max (max u3 u1) u2) u3} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E))) (Set.instMembershipSet.{max (max u3 u2) u1} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E))) e' (FiberPrebundle.pretrivializationAtlas.{u3, u2, u1} B F E _inst_1 _inst_2 a)) -> (IsOpen.{max u3 u2} (Prod.{u3, u2} B F) (instTopologicalSpaceProd.{u3, u2} B F _inst_1 _inst_2) (Inter.inter.{max u3 u2} (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instInterSet.{max u3 u2} (Prod.{u3, u2} B F)) (LocalEquiv.target.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E) e')) (Set.preimage.{max u3 u2, max u3 u1} (Prod.{u3, u2} B F) (Bundle.TotalSpace.{u3, u1} B E) (LocalEquiv.toFun.{max u3 u2, max u3 u1} (Prod.{u3, u2} B F) (Bundle.TotalSpace.{u3, u1} B E) (LocalEquiv.symm.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E) e'))) (LocalEquiv.source.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u1} B E) e))))) + forall {B : Type.{u3}} {F : Type.{u2}} {E : B -> Type.{u1}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u3, u2, u1} B F E _inst_2 _inst_3) (e : Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E)) (e' : Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E)), (Membership.mem.{max (max u3 u2) u1, max (max u3 u2) u1} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E)) (Set.{max (max (max u3 u1) u2) u3} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E))) (Set.instMembershipSet.{max (max u3 u2) u1} (Pretrivialization.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E))) e' (FiberPrebundle.pretrivializationAtlas.{u3, u2, u1} B F E _inst_2 _inst_3 a)) -> (IsOpen.{max u3 u2} (Prod.{u3, u2} B F) (instTopologicalSpaceProd.{u3, u2} B F _inst_2 _inst_3) (Inter.inter.{max u3 u2} (Set.{max u3 u2} (Prod.{u3, u2} B F)) (Set.instInterSet.{max u3 u2} (Prod.{u3, u2} B F)) (LocalEquiv.target.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E) e')) (Set.preimage.{max u3 u2, max u3 u1} (Prod.{u3, u2} B F) (Bundle.TotalSpace.{u3, u1} B E) (LocalEquiv.toFun.{max u3 u2, max u3 u1} (Prod.{u3, u2} B F) (Bundle.TotalSpace.{u3, u1} B E) (LocalEquiv.symm.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E) e'))) (LocalEquiv.source.{max u3 u1, max u3 u2} (Bundle.TotalSpace.{u3, u1} B E) (Prod.{u3, u2} B F) (Pretrivialization.toLocalEquiv.{u3, u2, max u3 u1} B F (Bundle.TotalSpace.{u3, u1} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u1} B E) e))))) Case conversion may be inaccurate. Consider using '#align fiber_prebundle.is_open_target_of_mem_pretrivialization_atlas_inter FiberPrebundle.isOpen_target_of_mem_pretrivializationAtlas_interₓ'. -/ theorem isOpen_target_of_mem_pretrivializationAtlas_inter (e e' : Pretrivialization F (π E)) (he' : e' ∈ a.pretrivializationAtlas) : @@ -1134,9 +1184,9 @@ def trivializationOfMemPretrivializationAtlas (he : e ∈ a.pretrivializationAtl /- warning: fiber_prebundle.mem_trivialization_at_source -> FiberPrebundle.mem_pretrivializationAt_source is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_1 _inst_2) (b : B) (x : E b), Membership.Mem.{max u1 u3, max u1 u3} (Bundle.TotalSpace.{u1, u3} B (fun (b : B) => E b)) (Set.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)) (Set.hasMem.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)) (Bundle.totalSpaceMk.{u1, u3} B (fun (b : B) => E b) b x) (LocalEquiv.source.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) (FiberPrebundle.pretrivializationAt.{u1, u2, u3} B F E _inst_1 _inst_2 a b))) + forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_2 _inst_3) (b : B) (x : E b), Membership.Mem.{max u1 u3, max u1 u3} (Bundle.TotalSpace.{u1, u3} B (fun (b : B) => E b)) (Set.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)) (Set.hasMem.{max u1 u3} (Bundle.TotalSpace.{u1, u3} B E)) (Bundle.totalSpaceMk.{u1, u3} B (fun (b : B) => E b) b x) (LocalEquiv.source.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) (FiberPrebundle.pretrivializationAt.{u1, u2, u3} B F E _inst_2 _inst_3 a b))) but is expected to have type - forall {B : Type.{u2}} {F : Type.{u1}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u2, u1, u3} B F E _inst_1 _inst_2) (b : B) (x : E b), Membership.mem.{max u3 u2, max u2 u3} (Bundle.TotalSpace.{u2, u3} B E) (Set.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B E)) (Set.instMembershipSet.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B E)) (Bundle.totalSpaceMk.{u2, u3} B E b x) (LocalEquiv.source.{max u2 u3, max u2 u1} (Bundle.TotalSpace.{u2, u3} B E) (Prod.{u2, u1} B F) (Pretrivialization.toLocalEquiv.{u2, u1, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u2, u3} B E) (FiberPrebundle.pretrivializationAt.{u2, u1, u3} B F E _inst_1 _inst_2 a b))) + forall {B : Type.{u2}} {F : Type.{u1}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u2, u1, u3} B F E _inst_2 _inst_3) (b : B) (x : E b), Membership.mem.{max u3 u2, max u2 u3} (Bundle.TotalSpace.{u2, u3} B E) (Set.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B E)) (Set.instMembershipSet.{max u2 u3} (Bundle.TotalSpace.{u2, u3} B E)) (Bundle.totalSpaceMk.{u2, u3} B E b x) (LocalEquiv.source.{max u2 u3, max u2 u1} (Bundle.TotalSpace.{u2, u3} B E) (Prod.{u2, u1} B F) (Pretrivialization.toLocalEquiv.{u2, u1, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u2, u3} B E) (FiberPrebundle.pretrivializationAt.{u2, u1, u3} B F E _inst_2 _inst_3 a b))) Case conversion may be inaccurate. Consider using '#align fiber_prebundle.mem_trivialization_at_source FiberPrebundle.mem_pretrivializationAt_sourceₓ'. -/ theorem mem_pretrivializationAt_source (b : B) (x : E b) : totalSpaceMk b x ∈ (a.pretrivializationAt b).source := @@ -1147,9 +1197,9 @@ theorem mem_pretrivializationAt_source (b : B) (x : E b) : /- warning: fiber_prebundle.total_space_mk_preimage_source -> FiberPrebundle.totalSpaceMk_preimage_source is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_1 _inst_2) (b : B), Eq.{succ u3} (Set.{u3} (E b)) (Set.preimage.{u3, max u1 u3} (E b) (Bundle.TotalSpace.{u1, u3} B E) (Bundle.totalSpaceMk.{u1, u3} B E b) (LocalEquiv.source.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) (FiberPrebundle.pretrivializationAt.{u1, u2, u3} B F E _inst_1 _inst_2 a b)))) (Set.univ.{u3} (E b)) + forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_2 _inst_3) (b : B), Eq.{succ u3} (Set.{u3} (E b)) (Set.preimage.{u3, max u1 u3} (E b) (Bundle.TotalSpace.{u1, u3} B E) (Bundle.totalSpaceMk.{u1, u3} B E b) (LocalEquiv.source.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) (FiberPrebundle.pretrivializationAt.{u1, u2, u3} B F E _inst_2 _inst_3 a b)))) (Set.univ.{u3} (E b)) but is expected to have type - forall {B : Type.{u2}} {F : Type.{u1}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u2, u1, u3} B F E _inst_1 _inst_2) (b : B), Eq.{succ u3} (Set.{u3} (E b)) (Set.preimage.{u3, max u2 u3} (E b) (Bundle.TotalSpace.{u2, u3} B E) (Bundle.totalSpaceMk.{u2, u3} B E b) (LocalEquiv.source.{max u2 u3, max u2 u1} (Bundle.TotalSpace.{u2, u3} B E) (Prod.{u2, u1} B F) (Pretrivialization.toLocalEquiv.{u2, u1, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u2, u3} B E) (FiberPrebundle.pretrivializationAt.{u2, u1, u3} B F E _inst_1 _inst_2 a b)))) (Set.univ.{u3} (E b)) + forall {B : Type.{u2}} {F : Type.{u1}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u2, u1, u3} B F E _inst_2 _inst_3) (b : B), Eq.{succ u3} (Set.{u3} (E b)) (Set.preimage.{u3, max u2 u3} (E b) (Bundle.TotalSpace.{u2, u3} B E) (Bundle.totalSpaceMk.{u2, u3} B E b) (LocalEquiv.source.{max u2 u3, max u2 u1} (Bundle.TotalSpace.{u2, u3} B E) (Prod.{u2, u1} B F) (Pretrivialization.toLocalEquiv.{u2, u1, max u2 u3} B F (Bundle.TotalSpace.{u2, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u2, u3} B E) (FiberPrebundle.pretrivializationAt.{u2, u1, u3} B F E _inst_2 _inst_3 a b)))) (Set.univ.{u3} (E b)) Case conversion may be inaccurate. Consider using '#align fiber_prebundle.total_space_mk_preimage_source FiberPrebundle.totalSpaceMk_preimage_sourceₓ'. -/ @[simp] theorem totalSpaceMk_preimage_source (b : B) : @@ -1171,9 +1221,9 @@ def fiberTopology (b : B) : TopologicalSpace (E b) := /- warning: fiber_prebundle.inducing_total_space_mk -> FiberPrebundle.inducing_totalSpaceMk is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_1 _inst_2) (b : B), Inducing.{u3, max u1 u3} (E b) (Bundle.TotalSpace.{u1, u3} B E) (FiberPrebundle.fiberTopology.{u1, u2, u3} B F E _inst_1 _inst_2 a b) (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_1 _inst_2 a) (Bundle.totalSpaceMk.{u1, u3} B E b) + forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_2 _inst_3) (b : B), Inducing.{u3, max u1 u3} (E b) (Bundle.TotalSpace.{u1, u3} B E) (FiberPrebundle.fiberTopology.{u1, u2, u3} B F E _inst_2 _inst_3 a b) (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_2 _inst_3 a) (Bundle.totalSpaceMk.{u1, u3} B E b) but is expected to have type - forall {B : Type.{u2}} {F : Type.{u1}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u2, u1, u3} B F E _inst_1 _inst_2) (b : B), Inducing.{u3, max u2 u3} (E b) (Bundle.TotalSpace.{u2, u3} B E) (FiberPrebundle.fiberTopology.{u2, u1, u3} B F E _inst_1 _inst_2 a b) (FiberPrebundle.totalSpaceTopology.{u2, u1, u3} B F E _inst_1 _inst_2 a) (Bundle.totalSpaceMk.{u2, u3} B E b) + forall {B : Type.{u2}} {F : Type.{u1}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u2, u1, u3} B F E _inst_2 _inst_3) (b : B), Inducing.{u3, max u2 u3} (E b) (Bundle.TotalSpace.{u2, u3} B E) (FiberPrebundle.fiberTopology.{u2, u1, u3} B F E _inst_2 _inst_3 a b) (FiberPrebundle.totalSpaceTopology.{u2, u1, u3} B F E _inst_2 _inst_3 a) (Bundle.totalSpaceMk.{u2, u3} B E b) Case conversion may be inaccurate. Consider using '#align fiber_prebundle.inducing_total_space_mk FiberPrebundle.inducing_totalSpaceMkₓ'. -/ @[continuity] theorem inducing_totalSpaceMk (b : B) : @@ -1186,9 +1236,9 @@ theorem inducing_totalSpaceMk (b : B) : /- warning: fiber_prebundle.continuous_total_space_mk -> FiberPrebundle.continuous_totalSpaceMk is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_1 _inst_2) (b : B), Continuous.{u3, max u1 u3} (E b) (Bundle.TotalSpace.{u1, u3} B E) (FiberPrebundle.fiberTopology.{u1, u2, u3} B F E _inst_1 _inst_2 a b) (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_1 _inst_2 a) (Bundle.totalSpaceMk.{u1, u3} B E b) + forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_2 _inst_3) (b : B), Continuous.{u3, max u1 u3} (E b) (Bundle.TotalSpace.{u1, u3} B E) (FiberPrebundle.fiberTopology.{u1, u2, u3} B F E _inst_2 _inst_3 a b) (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_2 _inst_3 a) (Bundle.totalSpaceMk.{u1, u3} B E b) but is expected to have type - forall {B : Type.{u2}} {F : Type.{u1}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u2} B] [_inst_2 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u2, u1, u3} B F E _inst_1 _inst_2) (b : B), Continuous.{u3, max u2 u3} (E b) (Bundle.TotalSpace.{u2, u3} B E) (FiberPrebundle.fiberTopology.{u2, u1, u3} B F E _inst_1 _inst_2 a b) (FiberPrebundle.totalSpaceTopology.{u2, u1, u3} B F E _inst_1 _inst_2 a) (Bundle.totalSpaceMk.{u2, u3} B E b) + forall {B : Type.{u2}} {F : Type.{u1}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u2} B] [_inst_3 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u2, u1, u3} B F E _inst_2 _inst_3) (b : B), Continuous.{u3, max u2 u3} (E b) (Bundle.TotalSpace.{u2, u3} B E) (FiberPrebundle.fiberTopology.{u2, u1, u3} B F E _inst_2 _inst_3 a b) (FiberPrebundle.totalSpaceTopology.{u2, u1, u3} B F E _inst_2 _inst_3 a) (Bundle.totalSpaceMk.{u2, u3} B E b) Case conversion may be inaccurate. Consider using '#align fiber_prebundle.continuous_total_space_mk FiberPrebundle.continuous_totalSpaceMkₓ'. -/ @[continuity] theorem continuous_totalSpaceMk (b : B) : @@ -1221,9 +1271,9 @@ def toFiberBundle : @FiberBundle B F _ _ E a.totalSpaceTopology a.fiberTopology /- warning: fiber_prebundle.continuous_proj -> FiberPrebundle.continuous_proj is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_1 _inst_2), Continuous.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_1 _inst_2 a) _inst_1 (Bundle.TotalSpace.proj.{u1, u3} B E) + forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_2 _inst_3), Continuous.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_2 _inst_3 a) _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) but is expected to have type - forall {B : Type.{u3}} {F : Type.{u1}} {E : B -> Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u3, u1, u2} B F E _inst_1 _inst_2), Continuous.{max u3 u2, u3} (Bundle.TotalSpace.{u3, u2} B E) B (FiberPrebundle.totalSpaceTopology.{u3, u1, u2} B F E _inst_1 _inst_2 a) _inst_1 (Bundle.TotalSpace.proj.{u3, u2} B E) + forall {B : Type.{u3}} {F : Type.{u1}} {E : B -> Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u3, u1, u2} B F E _inst_2 _inst_3), Continuous.{max u3 u2, u3} (Bundle.TotalSpace.{u3, u2} B E) B (FiberPrebundle.totalSpaceTopology.{u3, u1, u2} B F E _inst_2 _inst_3 a) _inst_2 (Bundle.TotalSpace.proj.{u3, u2} B E) Case conversion may be inaccurate. Consider using '#align fiber_prebundle.continuous_proj FiberPrebundle.continuous_projₓ'. -/ theorem continuous_proj : @Continuous _ _ a.totalSpaceTopology _ (π E) := by @@ -1235,9 +1285,9 @@ theorem continuous_proj : @Continuous _ _ a.totalSpaceTopology _ (π E) := /- warning: fiber_prebundle.continuous_on_of_comp_right -> FiberPrebundle.continuousOn_of_comp_right is a dubious translation: lean 3 declaration is - forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_1 : TopologicalSpace.{u1} B] [_inst_2 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_1 _inst_2) {X : Type.{u4}} [_inst_3 : TopologicalSpace.{u4} X] {f : (Bundle.TotalSpace.{u1, u3} B E) -> X} {s : Set.{u1} B}, (IsOpen.{u1} B _inst_1 s) -> (forall (b : B), (Membership.Mem.{u1, u1} B (Set.{u1} B) (Set.hasMem.{u1} B) b s) -> (ContinuousOn.{max u1 u2, u4} (Prod.{u1, u2} B F) X (Prod.topologicalSpace.{u1, u2} B F _inst_1 _inst_2) _inst_3 (Function.comp.{succ (max u1 u2), max (succ u1) (succ u3), succ u4} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E) X f (coeFn.{max (succ (max u1 u2)) (succ (max u1 u3)), max (succ (max u1 u2)) (succ (max u1 u3))} (LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (fun (_x : LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) => (Prod.{u1, u2} B F) -> (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.hasCoeToFun.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.symm.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) (FiberPrebundle.pretrivializationAt.{u1, u2, u3} B F E _inst_1 _inst_2 a b))))) (Set.prod.{u1, u2} B F (Inter.inter.{u1} (Set.{u1} B) (Set.hasInter.{u1} B) s (Pretrivialization.baseSet.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u1, u3} B E) (FiberPrebundle.pretrivializationAt.{u1, u2, u3} B F E _inst_1 _inst_2 a b))) (Set.univ.{u2} F)))) -> (ContinuousOn.{max u1 u3, u4} (Bundle.TotalSpace.{u1, u3} B E) X (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_1 _inst_2 a) _inst_3 f (Set.preimage.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B (Bundle.TotalSpace.proj.{u1, u3} B E) s)) + forall {B : Type.{u1}} {F : Type.{u2}} {E : B -> Type.{u3}} [_inst_2 : TopologicalSpace.{u1} B] [_inst_3 : TopologicalSpace.{u2} F] (a : FiberPrebundle.{u1, u2, u3} B F E _inst_2 _inst_3) {X : Type.{u4}} [_inst_4 : TopologicalSpace.{u4} X] {f : (Bundle.TotalSpace.{u1, u3} B E) -> X} {s : Set.{u1} B}, (IsOpen.{u1} B _inst_2 s) -> (forall (b : B), (Membership.Mem.{u1, u1} B (Set.{u1} B) (Set.hasMem.{u1} B) b s) -> (ContinuousOn.{max u1 u2, u4} (Prod.{u1, u2} B F) X (Prod.topologicalSpace.{u1, u2} B F _inst_2 _inst_3) _inst_4 (Function.comp.{succ (max u1 u2), max (succ u1) (succ u3), succ u4} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E) X f (coeFn.{max (succ (max u1 u2)) (succ (max u1 u3)), max (succ (max u1 u2)) (succ (max u1 u3))} (LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (fun (_x : LocalEquiv.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) => (Prod.{u1, u2} B F) -> (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.hasCoeToFun.{max u1 u2, max u1 u3} (Prod.{u1, u2} B F) (Bundle.TotalSpace.{u1, u3} B E)) (LocalEquiv.symm.{max u1 u3, max u1 u2} (Bundle.TotalSpace.{u1, u3} B E) (Prod.{u1, u2} B F) (Pretrivialization.toLocalEquiv.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) (FiberPrebundle.pretrivializationAt.{u1, u2, u3} B F E _inst_2 _inst_3 a b))))) (Set.prod.{u1, u2} B F (Inter.inter.{u1} (Set.{u1} B) (Set.hasInter.{u1} B) s (Pretrivialization.baseSet.{u1, u2, max u1 u3} B F (Bundle.TotalSpace.{u1, u3} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u1, u3} B E) (FiberPrebundle.pretrivializationAt.{u1, u2, u3} B F E _inst_2 _inst_3 a b))) (Set.univ.{u2} F)))) -> (ContinuousOn.{max u1 u3, u4} (Bundle.TotalSpace.{u1, u3} B E) X (FiberPrebundle.totalSpaceTopology.{u1, u2, u3} B F E _inst_2 _inst_3 a) _inst_4 f (Set.preimage.{max u1 u3, u1} (Bundle.TotalSpace.{u1, u3} B E) B (Bundle.TotalSpace.proj.{u1, u3} B E) s)) but is expected to have type - forall {B : Type.{u3}} {F : Type.{u1}} {E : B -> Type.{u2}} [_inst_1 : TopologicalSpace.{u3} B] [_inst_2 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u3, u1, u2} B F E _inst_1 _inst_2) {X : Type.{u4}} [_inst_3 : TopologicalSpace.{u4} X] {f : (Bundle.TotalSpace.{u3, u2} B E) -> X} {s : Set.{u3} B}, (IsOpen.{u3} B _inst_1 s) -> (forall (b : B), (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) b s) -> (ContinuousOn.{max u3 u1, u4} (Prod.{u3, u1} B F) X (instTopologicalSpaceProd.{u3, u1} B F _inst_1 _inst_2) _inst_3 (Function.comp.{succ (max u3 u1), max (succ u3) (succ u2), succ u4} (Prod.{u3, u1} B F) (Bundle.TotalSpace.{u3, u2} B E) X f (LocalEquiv.toFun.{max u3 u1, max u3 u2} (Prod.{u3, u1} B F) (Bundle.TotalSpace.{u3, u2} B E) (LocalEquiv.symm.{max u3 u2, max u3 u1} (Bundle.TotalSpace.{u3, u2} B E) (Prod.{u3, u1} B F) (Pretrivialization.toLocalEquiv.{u3, u1, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u2} B E) (FiberPrebundle.pretrivializationAt.{u3, u1, u2} B F E _inst_1 _inst_2 a b))))) (Set.prod.{u3, u1} B F (Inter.inter.{u3} (Set.{u3} B) (Set.instInterSet.{u3} B) s (Pretrivialization.baseSet.{u3, u1, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B E) _inst_1 _inst_2 (Bundle.TotalSpace.proj.{u3, u2} B E) (FiberPrebundle.pretrivializationAt.{u3, u1, u2} B F E _inst_1 _inst_2 a b))) (Set.univ.{u1} F)))) -> (ContinuousOn.{max u3 u2, u4} (Bundle.TotalSpace.{u3, u2} B E) X (FiberPrebundle.totalSpaceTopology.{u3, u1, u2} B F E _inst_1 _inst_2 a) _inst_3 f (Set.preimage.{max u3 u2, u3} (Bundle.TotalSpace.{u3, u2} B E) B (Bundle.TotalSpace.proj.{u3, u2} B E) s)) + forall {B : Type.{u3}} {F : Type.{u1}} {E : B -> Type.{u2}} [_inst_2 : TopologicalSpace.{u3} B] [_inst_3 : TopologicalSpace.{u1} F] (a : FiberPrebundle.{u3, u1, u2} B F E _inst_2 _inst_3) {X : Type.{u4}} [_inst_4 : TopologicalSpace.{u4} X] {f : (Bundle.TotalSpace.{u3, u2} B E) -> X} {s : Set.{u3} B}, (IsOpen.{u3} B _inst_2 s) -> (forall (b : B), (Membership.mem.{u3, u3} B (Set.{u3} B) (Set.instMembershipSet.{u3} B) b s) -> (ContinuousOn.{max u3 u1, u4} (Prod.{u3, u1} B F) X (instTopologicalSpaceProd.{u3, u1} B F _inst_2 _inst_3) _inst_4 (Function.comp.{succ (max u3 u1), max (succ u3) (succ u2), succ u4} (Prod.{u3, u1} B F) (Bundle.TotalSpace.{u3, u2} B E) X f (LocalEquiv.toFun.{max u3 u1, max u3 u2} (Prod.{u3, u1} B F) (Bundle.TotalSpace.{u3, u2} B E) (LocalEquiv.symm.{max u3 u2, max u3 u1} (Bundle.TotalSpace.{u3, u2} B E) (Prod.{u3, u1} B F) (Pretrivialization.toLocalEquiv.{u3, u1, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u2} B E) (FiberPrebundle.pretrivializationAt.{u3, u1, u2} B F E _inst_2 _inst_3 a b))))) (Set.prod.{u3, u1} B F (Inter.inter.{u3} (Set.{u3} B) (Set.instInterSet.{u3} B) s (Pretrivialization.baseSet.{u3, u1, max u3 u2} B F (Bundle.TotalSpace.{u3, u2} B E) _inst_2 _inst_3 (Bundle.TotalSpace.proj.{u3, u2} B E) (FiberPrebundle.pretrivializationAt.{u3, u1, u2} B F E _inst_2 _inst_3 a b))) (Set.univ.{u1} F)))) -> (ContinuousOn.{max u3 u2, u4} (Bundle.TotalSpace.{u3, u2} B E) X (FiberPrebundle.totalSpaceTopology.{u3, u1, u2} B F E _inst_2 _inst_3 a) _inst_4 f (Set.preimage.{max u3 u2, u3} (Bundle.TotalSpace.{u3, u2} B E) B (Bundle.TotalSpace.proj.{u3, u2} B E) s)) Case conversion may be inaccurate. Consider using '#align fiber_prebundle.continuous_on_of_comp_right FiberPrebundle.continuousOn_of_comp_rightₓ'. -/ /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ /-- For a fiber bundle `E` over `B` constructed using the `fiber_prebundle` mechanism, diff --git a/Mathbin/Topology/VectorBundle/Basic.lean b/Mathbin/Topology/VectorBundle/Basic.lean index 8aa86332d7..01d9fbd679 100644 --- a/Mathbin/Topology/VectorBundle/Basic.lean +++ b/Mathbin/Topology/VectorBundle/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri, Sebastien Gouezel, Heather Macbeth, Patrick Massot, Floris van Doorn ! This file was ported from Lean 3 source module topology.vector_bundle.basic -! leanprover-community/mathlib commit be2c24f56783935652cefffb4bfca7e4b25d167e +! leanprover-community/mathlib commit 0187644979f2d3e10a06e916a869c994facd9a87 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -361,6 +361,26 @@ end TopologicalVectorSpace section +namespace Bundle + +/-- The zero section of a vector bundle -/ +def zeroSection [∀ x, Zero (E x)] : B → TotalSpace E := fun x => totalSpaceMk x 0 +#align bundle.zero_section Bundle.zeroSection + +@[simp, mfld_simps] +theorem zeroSection_proj [∀ x, Zero (E x)] (x : B) : (zeroSection E x).proj = x := + rfl +#align bundle.zero_section_proj Bundle.zeroSection_proj + +@[simp, mfld_simps] +theorem zeroSection_snd [∀ x, Zero (E x)] (x : B) : (zeroSection E x).2 = 0 := + rfl +#align bundle.zero_section_snd Bundle.zeroSection_snd + +end Bundle + +open Bundle + variable [NontriviallyNormedField R] [∀ x, AddCommMonoid (E x)] [∀ x, Module R (E x)] [NormedAddCommGroup F] [NormedSpace R F] [TopologicalSpace B] [TopologicalSpace (TotalSpace E)] [∀ x, TopologicalSpace (E x)] [FiberBundle F E] @@ -408,8 +428,8 @@ def continuousLinearMapAt (e : Trivialization F (π E)) [e.isLinear R] (b : B) : refine' continuous_if_const _ (fun hb => _) fun _ => continuous_zero exact continuous_snd.comp - (e.to_local_homeomorph.continuous_on.comp_continuous - (FiberBundle.totalSpaceMk_inducing F E b).Continuous fun x => e.mem_source.mpr hb) } + (e.continuous_on.comp_continuous (FiberBundle.totalSpaceMk_inducing F E b).Continuous + fun x => e.mem_source.mpr hb) } #align trivialization.continuous_linear_map_at Trivialization.continuousLinearMapAt /-- Backwards map of `continuous_linear_equiv_at`, defined everywhere. -/ @@ -455,8 +475,8 @@ def continuousLinearEquivAt (e : Trivialization F (π E)) [e.isLinear R] (b : B) invFun := e.symm b continuous_toFun := continuous_snd.comp - (e.toLocalHomeomorph.ContinuousOn.comp_continuous - (FiberBundle.totalSpaceMk_inducing F E b).Continuous fun x => e.mem_source.mpr hb) + (e.ContinuousOn.comp_continuous (FiberBundle.totalSpaceMk_inducing F E b).Continuous + fun x => e.mem_source.mpr hb) continuous_invFun := (e.symmL R b).Continuous } #align trivialization.continuous_linear_equiv_at Trivialization.continuousLinearEquivAt @@ -485,8 +505,7 @@ theorem continuousLinearEquivAt_apply' (e : Trivialization F (π E)) [e.isLinear variable (R) theorem apply_eq_prod_continuousLinearEquivAt (e : Trivialization F (π E)) [e.isLinear R] (b : B) - (hb : b ∈ e.baseSet) (z : E b) : - e.toLocalHomeomorph ⟨b, z⟩ = (b, e.continuousLinearEquivAt R b hb z) := + (hb : b ∈ e.baseSet) (z : E b) : e ⟨b, z⟩ = (b, e.continuousLinearEquivAt R b hb z) := by ext · refine' e.coe_fst _ @@ -495,19 +514,24 @@ theorem apply_eq_prod_continuousLinearEquivAt (e : Trivialization F (π E)) [e.i · simp only [coe_coe, continuous_linear_equiv_at_apply] #align trivialization.apply_eq_prod_continuous_linear_equiv_at Trivialization.apply_eq_prod_continuousLinearEquivAt +protected theorem zeroSection (e : Trivialization F (π E)) [e.isLinear R] {x : B} + (hx : x ∈ e.baseSet) : e (zeroSection E x) = (x, 0) := by + simp_rw [zero_section, total_space_mk, e.apply_eq_prod_continuous_linear_equiv_at R x hx 0, + map_zero] +#align trivialization.zero_section Trivialization.zeroSection + variable {R} theorem symm_apply_eq_mk_continuousLinearEquivAt_symm (e : Trivialization F (π E)) [e.isLinear R] (b : B) (hb : b ∈ e.baseSet) (z : F) : e.toLocalHomeomorph.symm ⟨b, z⟩ = totalSpaceMk b ((e.continuousLinearEquivAt R b hb).symm z) := by - have h : (b, z) ∈ e.to_local_homeomorph.target := - by + have h : (b, z) ∈ e.target := by rw [e.target_eq] exact ⟨hb, mem_univ _⟩ - apply e.to_local_homeomorph.inj_on (e.to_local_homeomorph.map_target h) + apply e.inj_on (e.map_target h) · simp only [e.source_eq, hb, mem_preimage] - simp_rw [e.apply_eq_prod_continuous_linear_equiv_at R b hb, e.to_local_homeomorph.right_inv h, + simp_rw [e.right_inv h, coe_coe, e.apply_eq_prod_continuous_linear_equiv_at R b hb, ContinuousLinearEquiv.apply_symm_apply] #align trivialization.symm_apply_eq_mk_continuous_linear_equiv_at_symm Trivialization.symm_apply_eq_mk_continuousLinearEquivAt_symm diff --git a/README.md b/README.md index c346324502..67f66e06a4 100644 --- a/README.md +++ b/README.md @@ -1,4 +1,4 @@ -Tracking mathlib commit: [`0148d455199ed64bf8eb2f493a1e7eb9211ce170`](https://github.com/leanprover-community/mathlib/commit/0148d455199ed64bf8eb2f493a1e7eb9211ce170) +Tracking mathlib commit: [`b685f506164f8d17a6404048bc4d696739c5d976`](https://github.com/leanprover-community/mathlib/commit/b685f506164f8d17a6404048bc4d696739c5d976) You should use this repository to inspect the Lean 4 files that `mathport` has generated from mathlib3. Please run `lake build` first, to download a copy of the generated `.olean` files. diff --git a/lake-manifest.json b/lake-manifest.json index 0471e47a21..517f7fc2d6 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": "5cb47ce226bc5e0380bbc2405bf70b4cce6529ce", + "rev": "b7c1f89544904b0e18c25a976554a607d5db4445", "name": "lean3port", - "inputRev?": "5cb47ce226bc5e0380bbc2405bf70b4cce6529ce"}}, + "inputRev?": "b7c1f89544904b0e18c25a976554a607d5db4445"}}, {"git": {"url": "https://github.com/leanprover-community/mathlib4.git", "subDir?": null, - "rev": "26daa9a9f2d032819a5c93099797dfafc8ed907a", + "rev": "7c82cee42662b7f41701b08a1b8aaaf6e6b19b36", "name": "mathlib", - "inputRev?": "26daa9a9f2d032819a5c93099797dfafc8ed907a"}}, + "inputRev?": "7c82cee42662b7f41701b08a1b8aaaf6e6b19b36"}}, {"git": {"url": "https://github.com/gebner/quote4", "subDir?": null, diff --git a/lakefile.lean b/lakefile.lean index d82da8a7a6..c5754ac29f 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-03-30-00" +def tag : String := "nightly-2023-03-30-02" 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"@"5cb47ce226bc5e0380bbc2405bf70b4cce6529ce" +require lean3port from git "https://github.com/leanprover-community/lean3port.git"@"b7c1f89544904b0e18c25a976554a607d5db4445" @[default_target] lean_lib Mathbin where diff --git a/upstream-rev b/upstream-rev index 531052e223..894548d528 100644 --- a/upstream-rev +++ b/upstream-rev @@ -1 +1 @@ -0148d455199ed64bf8eb2f493a1e7eb9211ce170 +b685f506164f8d17a6404048bc4d696739c5d976